@[inline]
Equations
- Except.map f x = match x with | Except.error err => Except.error err | Except.ok v => Except.ok (f v)
Instances For
@[inline]
Equations
- Except.mapError f x = match x with | Except.error err => Except.error (f err) | Except.ok v => Except.ok v
Instances For
@[inline]
def
Except.bind
{ε : Type u}
{α : Type u_1}
{β : Type u_2}
(ma : Except ε α)
(f : α → Except ε β)
 :
Except ε β
Equations
- ma.bind f = match ma with | Except.error err => Except.error err | Except.ok v => f v
Instances For
@[inline]
Equations
- x.toOption = match x with | Except.ok a => some a | Except.error a => none
Instances For
@[inline]
def
ExceptT.bindCont
{ε : Type u}
{m : Type u → Type v}
[Monad m]
{α : Type u}
{β : Type u}
(f : α → ExceptT ε m β)
 :
Equations
- ExceptT.bindCont f x = match x with | Except.ok a => f a | Except.error e => pure (Except.error e)
Instances For
@[inline]
def
ExceptT.bind
{ε : Type u}
{m : Type u → Type v}
[Monad m]
{α : Type u}
{β : Type u}
(ma : ExceptT ε m α)
(f : α → ExceptT ε m β)
 :
ExceptT ε m β
Equations
- ma.bind f = ExceptT.mk (ma >>= ExceptT.bindCont f)
Instances For
@[inline]
def
ExceptT.map
{ε : Type u}
{m : Type u → Type v}
[Monad m]
{α : Type u}
{β : Type u}
(f : α → β)
(x : ExceptT ε m α)
 :
ExceptT ε m β
Equations
- ExceptT.map f x = ExceptT.mk do let a ← x match a with | Except.ok a => pure (Except.ok (f a)) | Except.error e => pure (Except.error e)
Instances For
@[inline]
def
ExceptT.tryCatch
{ε : Type u}
{m : Type u → Type v}
[Monad m]
{α : Type u}
(ma : ExceptT ε m α)
(handle : ε → ExceptT ε m α)
 :
ExceptT ε m α
Equations
- ma.tryCatch handle = ExceptT.mk do let res ← ma match res with | Except.ok a => pure (Except.ok a) | Except.error e => handle e
Instances For
@[inline]
def
ExceptT.adapt
{ε : Type u}
{m : Type u → Type v}
[Monad m]
{ε' : Type u}
{α : Type u}
(f : ε → ε')
 :
Equations
- ExceptT.adapt f x = ExceptT.mk (Except.mapError f <$> x)
Instances For
@[always_inline]
instance
instMonadExceptOfExceptTOfMonad
(m : Type u → Type v)
(ε₁ : Type u)
(ε₂ : Type u)
[Monad m]
[MonadExceptOf ε₁ m]
 :
MonadExceptOf ε₁ (ExceptT ε₂ m)
Equations
- One or more equations did not get rendered due to their size.
@[always_inline]
instance
instMonadExceptOfExceptTOfMonad_1
(m : Type u → Type v)
(ε : Type u)
[Monad m]
 :
MonadExceptOf ε (ExceptT ε m)
Equations
- instMonadExceptOfExceptTOfMonad_1 m ε = { throw := fun {α : Type u} (e : ε) => ExceptT.mk (pure (Except.error e)), tryCatch := fun {α : Type u} => ExceptT.tryCatch }
Equations
- instMonadExceptOfExcept ε = { throw := fun {α : Type u_2} => Except.error, tryCatch := fun {α : Type u_2} => Except.tryCatch }
@[inline]
def
MonadExcept.orelse'
{ε : Type u}
{m : Type v → Type w}
[MonadExcept ε m]
{α : Type v}
(t₁ : m α)
(t₂ : m α)
(useFirstEx : optParam Bool true)
 :
m α
Alternative orelse operator that allows to select which exception should be used.
The default is to use the first exception since the standard orelse uses the second.
Equations
Instances For
def
liftExcept
{ε : Type u_1}
{m : Type u_2 → Type u_3}
{α : Type u_2}
[MonadExceptOf ε m]
[Pure m]
 :
Except ε α → m α
Equations
- liftExcept x = match x with | Except.ok a => pure a | Except.error e => throw e
Instances For
instance
instMonadControlExceptTOfMonad
(ε : Type u)
(m : Type u → Type v)
[Monad m]
 :
MonadControl m (ExceptT ε m)
Equations
- One or more equations did not get rendered due to their size.
- tryFinally' x fruns- xand then the "finally" computation- f. When- xsucceeds with- a : α,- f (some a)is returned. If- xfails for- m's definition of failure,- f noneis returned. Hence- tryFinally'can be thought of as performing the same role as a- finallyblock in an imperative programming language.
Instances
@[inline]
def
tryFinally
{m : Type u → Type v}
{α : Type u}
{β : Type u}
[MonadFinally m]
[Functor m]
(x : m α)
(finalizer : m β)
 :
m α
Execute x and then execute finalizer even if x threw an exception
Equations
- tryFinally x finalizer = let y := tryFinally' x fun (x : Option α) => finalizer; (fun (x : α × β) => x.fst) <$> y
Instances For
@[always_inline]
@[always_inline]
instance
ExceptT.finally
{m : Type u → Type v}
{ε : Type u}
[MonadFinally m]
[Monad m]
 :
MonadFinally (ExceptT ε m)
Equations
- One or more equations did not get rendered due to their size.