ExceptT #
@[simp]
theorem
ExceptT.run_lift
{m : Type u → Type v}
{α ε : Type u}
[Monad m]
(x : m α)
:
(ExceptT.lift x).run = Except.ok <$> x
@[simp]
theorem
ExceptT.run_throw
{m : Type u_1 → Type u_2}
{ε β : Type u_1}
{e : ε}
[Monad m]
:
(throw e).run = pure (Except.error e)
@[simp]
theorem
ExceptT.run_bind_lift
{m : Type u_1 → Type u_2}
{α ε β : Type u_1}
[Monad m]
[LawfulMonad m]
(x : m α)
(f : α → ExceptT ε m β)
:
(ExceptT.lift x >>= f).run = do
let a ← x
(f a).run
theorem
ExceptT.run_bind
{m : Type u_1 → Type u_2}
{ε α β : Type u_1}
{f : α → ExceptT ε m β}
[Monad m]
(x : ExceptT ε m α)
:
(x >>= f).run = do
let x ← x.run
match x with
| Except.ok x => (f x).run
| Except.error e => pure (Except.error e)
@[simp]
theorem
ExceptT.lift_pure
{m : Type u_1 → Type u_2}
{α ε : Type u_1}
[Monad m]
[LawfulMonad m]
(a : α)
:
ExceptT.lift (pure a) = pure a
@[simp]
theorem
ExceptT.run_map
{m : Type u_1 → Type u_2}
{α β ε : Type u_1}
[Monad m]
[LawfulMonad m]
(f : α → β)
(x : ExceptT ε m α)
:
(f <$> x).run = Except.map f <$> x.run
theorem
ExceptT.seqLeft_eq
{α β ε : Type u}
{m : Type u → Type v}
[Monad m]
[LawfulMonad m]
(x : ExceptT ε m α)
(y : ExceptT ε m β)
:
x <* y = Function.const β <$> x <*> y
theorem
ExceptT.seqRight_eq
{m : Type u_1 → Type u_2}
{ε α β : Type u_1}
[Monad m]
[LawfulMonad m]
(x : ExceptT ε m α)
(y : ExceptT ε m β)
:
x *> y = Function.const α id <$> x <*> y
instance
ExceptT.instLawfulMonad
{m : Type u_1 → Type u_2}
{ε : Type u_1}
[Monad m]
[LawfulMonad m]
:
LawfulMonad (ExceptT ε m)
Equations
- ⋯ = ⋯
Except #
Equations
- ⋯ = ⋯
ReaderT #
@[simp]
theorem
ReaderT.run_mapConst
{m : Type u_1 → Type u_2}
{α ρ β : Type u_1}
[Monad m]
(a : α)
(x : ReaderT ρ m β)
(ctx : ρ)
:
(Functor.mapConst a x).run ctx = Functor.mapConst a (x.run ctx)
instance
ReaderT.instLawfulFunctor
{m : Type u_1 → Type u_2}
{ρ : Type u_1}
[Monad m]
[LawfulFunctor m]
:
LawfulFunctor (ReaderT ρ m)
Equations
- ⋯ = ⋯
instance
ReaderT.instLawfulApplicative
{m : Type u_1 → Type u_2}
{ρ : Type u_1}
[Monad m]
[LawfulApplicative m]
:
LawfulApplicative (ReaderT ρ m)
Equations
- ⋯ = ⋯
instance
ReaderT.instLawfulMonad
{m : Type u_1 → Type u_2}
{ρ : Type u_1}
[Monad m]
[LawfulMonad m]
:
LawfulMonad (ReaderT ρ m)
Equations
- ⋯ = ⋯
StateRefT #
instance
instLawfulMonadStateRefT'
{m : Type → Type}
{ω σ : Type}
[Monad m]
[LawfulMonad m]
:
LawfulMonad (StateRefT' ω σ m)
Equations
- ⋯ = ⋯
StateT #
@[simp]
theorem
StateT.run_set
{m : Type u_1 → Type u_2}
{σ : Type u_1}
[Monad m]
(s s' : σ)
:
(set s').run s = pure (PUnit.unit, s')
@[simp]
theorem
StateT.run_modify
{m : Type u_1 → Type u_2}
{σ : Type u_1}
[Monad m]
(f : σ → σ)
(s : σ)
:
(modify f).run s = pure (PUnit.unit, f s)
@[simp]
theorem
StateT.run_lift
{m : Type u → Type u_1}
{α σ : Type u}
[Monad m]
(x : m α)
(s : σ)
:
(StateT.lift x).run s = do
let a ← x
pure (a, s)
theorem
StateT.run_bind_lift
{m : Type u → Type u_1}
{β α σ : Type u}
[Monad m]
[LawfulMonad m]
(x : m α)
(f : α → StateT σ m β)
(s : σ)
:
(StateT.lift x >>= f).run s = do
let a ← x
(f a).run s
theorem
StateT.seqRight_eq
{m : Type u_1 → Type u_2}
{σ α β : Type u_1}
[Monad m]
[LawfulMonad m]
(x : StateT σ m α)
(y : StateT σ m β)
:
x *> y = Function.const α id <$> x <*> y
theorem
StateT.seqLeft_eq
{m : Type u_1 → Type u_2}
{σ α β : Type u_1}
[Monad m]
[LawfulMonad m]
(x : StateT σ m α)
(y : StateT σ m β)
:
x <* y = Function.const β <$> x <*> y
instance
StateT.instLawfulMonad
{m : Type u_1 → Type u_2}
{σ : Type u_1}
[Monad m]
[LawfulMonad m]
:
LawfulMonad (StateT σ m)
Equations
- ⋯ = ⋯