The combined state and list monad transformer.
StateListT σ α
is equivalent to StateT σ (ListT α)
but more efficient.
WARNING: StateListT σ α m
is only a monad if m
is a commutative monad.
For example,
def problem : StateListT Unit (StateM (Array Nat)) Unit := do
Alternative.orElse (pure ()) (fun _ => pure ())
StateListT.lift $ modify (·.push 0)
StateListT.lift $ modify (·.push 1)
#eval ((problem.run' ()).run #[]).2
will yield either #[0,1,0,1]
, or #[0,0,1,1]
, depending on the order in which the actions
in the do block are combined.
StateList
Equations
The combined state and list monad transformer.
Equations
- Mathlib.Meta.FunProp.StateListT σ m α = (σ → m (Mathlib.Meta.FunProp.StateList σ α))
Instances For
@[inline]
def
Mathlib.Meta.FunProp.StateListT.run
{α σ : Type u}
{m : Type u → Type v}
[Functor m]
(x : StateListT σ m α)
(s : σ)
:
Run x
on a given state s
, returning the list of values with corresponding states.
Equations
- x.run s = Mathlib.Meta.FunProp.StateList.toList✝ <$> x s
Instances For
@[inline]
def
Mathlib.Meta.FunProp.StateListT.run'
{α σ : Type u}
{m : Type u → Type v}
[Functor m]
(x : StateListT σ m α)
(s : σ)
:
m (List α)
Run x
on a given state s
, returning the list of values.
Equations
- x.run' s = Mathlib.Meta.FunProp.StateList.toList'✝ <$> x s
Instances For
@[reducible, inline]
The combined state and list monad.
Equations
Instances For
@[always_inline]
instance
Mathlib.Meta.FunProp.StateListT.instMonad
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
Monad (StateListT σ m)
instance
Mathlib.Meta.FunProp.StateListT.instAlternative
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
Alternative (StateListT σ m)
Equations
- One or more equations did not get rendered due to their size.
@[inline]
def
Mathlib.Meta.FunProp.StateListT.get
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
StateListT σ m σ
Return the state from StateListT σ m
.
Equations
Instances For
@[inline]
def
Mathlib.Meta.FunProp.StateListT.set
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
σ → StateListT σ m PUnit
Set the state in StateListT σ m
.
Equations
Instances For
@[inline]
def
Mathlib.Meta.FunProp.StateListT.modifyGet
{α σ : Type u}
{m : Type u → Type v}
[Monad m]
(f : σ → α × σ)
:
StateListT σ m α
Modify and get the state in StateListT σ m
.
Equations
- Mathlib.Meta.FunProp.StateListT.modifyGet f s = pure (Mathlib.Meta.FunProp.StateList.cons (f s).fst (f s).snd Mathlib.Meta.FunProp.StateList.nil)
Instances For
@[inline]
def
Mathlib.Meta.FunProp.StateListT.lift
{α σ : Type u}
{m : Type u → Type v}
[Monad m]
(t : m α)
:
StateListT σ m α
Lift an action from m α
to StateListT σ m α
.
Equations
Instances For
instance
Mathlib.Meta.FunProp.StateListT.instMonadLift
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
MonadLift m (StateListT σ m)
Equations
- Mathlib.Meta.FunProp.StateListT.instMonadLift = { monadLift := fun {α : Type ?u.26} => Mathlib.Meta.FunProp.StateListT.lift }
@[always_inline]
instance
Mathlib.Meta.FunProp.StateListT.instMonadFunctor
{σ : Type u}
{m : Type u → Type v}
:
MonadFunctor m (StateListT σ m)
Equations
- One or more equations did not get rendered due to their size.
@[always_inline]
instance
Mathlib.Meta.FunProp.StateListT.instMonadExceptOf
{σ : Type u}
{m : Type u → Type v}
[Monad m]
{ε : Type u_1}
[MonadExceptOf ε m]
:
MonadExceptOf ε (StateListT σ m)
Equations
- One or more equations did not get rendered due to their size.
instance
Mathlib.Meta.FunProp.instMonadStateOfStateListT
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
MonadStateOf σ (StateListT σ m)
Equations
- One or more equations did not get rendered due to their size.
@[always_inline]
instance
Mathlib.Meta.FunProp.StateListT.monadControl
{σ : Type u}
{m : Type u → Type v}
[Monad m]
:
MonadControl m (StateListT σ m)
Equations
- One or more equations did not get rendered due to their size.