doc-next-gen

Init.Control.Basic

Module docstring

{"# How MonadControl works

There is a tutorial by Alexis King that this docstring is based on.

Suppose we have foo : ∀ α, IO α → IO α and bar : StateT σ IO β (ie, bar : σ → IO (σ × β)). We might want to 'map' bar by foo. Concretely we would write this as:

```lean opaque foo : ∀ {α}, IO α → IO α opaque bar : StateT σ IO β

def mapped_foo : StateT σ IO β := do let s ← get let (b, s') ← liftM <| foo <| StateT.run bar s set s' return b ```

This is fine but it's not going to generalise, what if we replace StateT Nat IO with a large tower of monad transformers? We would have to rewrite the above to handle each of the run functions for each transformer in the stack.

Is there a way to generalise run as a kind of inverse of lift? We have lift : m α → StateT σ m α for all m, but we also need to 'unlift' the state. But unlift : StateT σ IO α → IO α can't be implemented. So we need something else.

If we look at the definition of mapped_foo, we see that lift <| foo <| StateT.run bar s has the type IO (σ × β). The key idea is that σ × β contains all of the information needed to reconstruct the state and the new value.

Now lets define some values to generalise mapped_foo: - Write IO (σ × β) as IO (stM β) - Write StateT.run . s as mapInBase : StateT σ IO α → IO (stM β) - Define restoreM : IO (stM α) → StateT σ IO α as below

```lean def stM (α : Type) := α × σ

def restoreM (x : IO (stM α)) : StateT σ IO α := do let (a,s) ← liftM x set s return a ```

To get:

lean def mapped_foo' : StateT σ IO β := do let s ← get let mapInBase := fun z => StateT.run z s restoreM <| foo <| mapInBase bar

and finally define

lean def control {α : Type} (f : ({β : Type} → StateT σ IO β → IO (stM β)) → IO (stM α)) : StateT σ IO α := do let s ← get let mapInBase := fun {β} (z : StateT σ IO β) => StateT.run z s let r : IO (stM α) := f mapInBase restoreM r

Now we can write mapped_foo as:

lean def mapped_foo'' : StateT σ IO β := control (fun mapInBase => foo (mapInBase bar))

The core idea of mapInBase is that given any β, it runs an instance of StateT σ IO β and 'packages' the result and state as IO (stM β) so that it can be piped through foo. Once it's been through foo we can then unpack the state again with restoreM. Hence we can apply foo to bar without losing track of the state.

Here stM β = σ × β is the 'packaged result state', but we can generalise: if we have a tower StateT σ₁ <| StateT σ₂ <| IO, then the composite packaged state is going to be stM₁₂ β := σ₁ × σ₂ × β or stM₁₂ := stM₁ ∘ stM₂.

MonadControl m n means that when programming in the monad n, we can switch to a base monad m using control, just like with liftM. In contrast to liftM, however, we also get a function runInBase that allows us to \"lower\" actions in n into m. This is really useful when we have large towers of monad transformers, as we do in the metaprogramming library.

For example there is a function withNewMCtxDepthImp : MetaM α → MetaM α that runs the input monad instance in a new nested metavariable context. We can lift this to withNewMctxDepth : n α → n α using MonadControlT MetaM n (MonadControlT is the transitive closure of MonadControl). Which means that we can also run withNewMctxDepth in the Tactic monad without needing to faff around with lifts and all the other boilerplate needed in mapped_foo.

Relationship to MonadFunctor

A stricter form of MonadControl is MonadFunctor, which defines monadMap {α} : (∀ {β}, m β → m β) → n α → n α. Using monadMap it is also possible to define mapped_foo above. However there are some mappings which can't be derived using MonadFunctor. For example:

```lean,ignore @[inline] def map1MetaM [MonadControlT MetaM n] [Monad n] (f : forall {α}, (β → MetaM α) → MetaM α) {α} (k : β → n α) : n α := control fun runInBase => f fun b => runInBase <| k b

@[inline] def map2MetaM [MonadControlT MetaM n] [Monad n] (f : forall {α}, (β → γ → MetaM α) → MetaM α) {α} (k : β → γ → n α) : n α := control fun runInBase => f fun b c => runInBase <| k b c ```

In monadMap, we can only 'run in base' a single computation in n into the base monad m. Using control means that runInBase can be used multiple times.

"}

instForInOfForIn' instance
[ForIn' m ρ α d] : ForIn m ρ α
Full source
/--
A `ForIn'` instance, which handles `for h : x in c do`,
can also handle `for x in x do` by ignoring `h`, and so provides a `ForIn` instance.

Note that this instance will cause a potentially non-defeq duplication if both `ForIn` and `ForIn'`
instances are provided for the same type.
-/
-- We set the priority to 500 so it is below the default,
-- but still above the low priority instance from `Stream`.
instance (priority := 500) instForInOfForIn' [ForIn' m ρ α d] : ForIn m ρ α where
  forIn x b f := forIn' x b fun a _ => f a
Standard Monadic Iteration from Proof-Carrying Iteration
Informal description
For any monad `m`, collection type `ρ`, and element type `α` with a membership relation, if there exists a `ForIn'` instance (supporting monadic iteration with membership proofs), then there also exists a `ForIn` instance (supporting standard monadic iteration) on the same types.
forIn'_eq_forIn theorem
[d : Membership α ρ] [ForIn' m ρ α d] {β} [Monad m] (x : ρ) (b : β) (f : (a : α) → a ∈ x → β → m (ForInStep β)) (g : (a : α) → β → m (ForInStep β)) (h : ∀ a m b, f a m b = g a b) : forIn' x b f = forIn x b g
Full source
@[simp] theorem forIn'_eq_forIn [d : Membership α ρ] [ForIn' m ρ α d] {β} [Monad m] (x : ρ) (b : β)
    (f : (a : α) → a ∈ x → β → m (ForInStep β)) (g : (a : α) → β → m (ForInStep β))
    (h : ∀ a m b, f a m b = g a b) :
    forIn' x b f = forIn x b g := by
  simp [instForInOfForIn']
  congr
  apply funext
  intro a
  apply funext
  intro m
  apply funext
  intro b
  simp [h]
  rfl
Equivalence of Monadic Iteration with and without Membership Proofs
Informal description
Let $m$ be a monad, $\rho$ a collection type with a membership relation for elements of type $\alpha$, and $\beta$ an arbitrary type. Given: 1. A collection $x \in \rho$, 2. An initial state $b \in \beta$, 3. Two functions $f : \alpha \to (a \in x) \to \beta \to m(\text{ForInStep}\ \beta)$ and $g : \alpha \to \beta \to m(\text{ForInStep}\ \beta)$, 4. The condition that for all $a \in \alpha$, membership proof $m$, and state $b$, we have $f\ a\ m\ b = g\ a\ b$. Then the monadic iteration with membership proofs `forIn'` is equal to the standard monadic iteration `forIn` when applied to these parameters.
forIn_eq_forin' theorem
[d : Membership α ρ] [ForIn' m ρ α d] {β} [Monad m] (x : ρ) (b : β) (f : (a : α) → β → m (ForInStep β)) : forIn x b f = forIn' x b (fun x h => binderNameHint x f <| binderNameHint h () <| f x)
Full source
@[wf_preprocess] theorem forIn_eq_forin' [d : Membership α ρ] [ForIn' m ρ α d] {β} [Monad m]
    (x : ρ) (b : β) (f : (a : α) → β → m (ForInStep β)) :
    forIn x b f = forIn' x b (fun x h => binderNameHint x f <| binderNameHint h () <| f x) := by
  simp [binderNameHint]
  rfl
Equivalence of Standard and Proof-Carrying Monadic Iteration
Informal description
Let $m$ be a monad, $\rho$ a collection type with a membership relation for elements of type $\alpha$, and $\beta$ an arbitrary type. Given: 1. A collection $x \in \rho$, 2. An initial state $b \in \beta$, 3. A function $f : \alpha \to \beta \to m(\text{ForInStep}\ \beta)$, then the standard monadic iteration `forIn` is equal to the monadic iteration with membership proofs `forIn'` when applied to these parameters, where the membership proof is handled via `binderNameHint`.
ForInStep.value definition
(x : ForInStep α) : α
Full source
/--
Extracts the value from a `ForInStep`, ignoring whether it is `ForInStep.done` or `ForInStep.yield`.
-/
def ForInStep.value (x : ForInStep α) : α :=
  match x with
  | ForInStep.done b => b
  | ForInStep.yield b => b
Extract value from loop control indicator
Informal description
Given a value `x` of type `ForInStep α`, which represents either an early termination (`done`) or continuation (`yield`) in a loop iteration, this function extracts the underlying value of type `α` regardless of the constructor used.
ForInStep.value_done theorem
(b : β) : (ForInStep.done b).value = b
Full source
@[simp] theorem ForInStep.value_done (b : β) : (ForInStep.done b).value = b := rfl
Value Extraction from Done Constructor in ForInStep
Informal description
For any value $b$ of type $\beta$, the value extracted from the `done` constructor of `ForInStep$ is equal to $b$, i.e., $(\text{ForInStep.done } b).\text{value} = b$.
ForInStep.value_yield theorem
(b : β) : (ForInStep.yield b).value = b
Full source
@[simp] theorem ForInStep.value_yield (b : β) : (ForInStep.yield b).value = b := rfl
Value Extraction from Yield Constructor in ForInStep
Informal description
For any value $b$ of type $\beta$, the value extracted from the `yield` constructor of `ForInStep` is equal to $b$, i.e., $(\text{ForInStep.yield } b).\text{value} = b$.
Functor.mapRev definition
{f : Type u → Type v} [Functor f] {α β : Type u} : f α → (α → β) → f β
Full source
/--
Maps a function over a functor, with parameters swapped so that the function comes last.

This function is `Functor.map` with the parameters reversed, typically used via the `<&>` operator.
-/
@[reducible]
def Functor.mapRev {f : Type u → Type v} [Functor f] {α β : Type u} : f α → (α → β) → f β :=
  fun a f => f <$> a
Reversed functor mapping
Informal description
Given a functor `f` and a function `g : α → β`, the function `mapRev` applies `g` to the contents of a functorial value `x : f α`, producing a new functorial value `f β`. This is equivalent to the standard `map` operation but with the arguments swapped, allowing the function to be written last. In symbols, for `x : f α` and `g : α → β`, we have: \[ \text{mapRev}(x, g) = g <\$> x \]
term_<&>_ definition
: Lean.TrailingParserDescr✝
Full source
@[inherit_doc Functor.mapRev]
infixr:100 " <&> " => Functor.mapRev
Reversed functor map operator `<&>`
Informal description
The infix operator `<&>` is defined as a reversed version of the functor mapping operation, where `x <&> f` applies the function `f` to the contents of the functorial value `x`. This is equivalent to `f <\$> x` but with the arguments swapped for better readability in certain contexts.
Functor.discard definition
{f : Type u → Type v} {α : Type u} [Functor f] (x : f α) : f PUnit
Full source
/--
Discards the value in a functor, retaining the functor's structure.

Discarding values is especially useful when using `Applicative` functors or `Monad`s to implement
effects, and some operation should be carried out only for its effects. In `do`-notation, statements
whose values are discarded must return `Unit`, and `discard` can be used to explicitly discard their
values.
-/
@[always_inline, inline]
def Functor.discard {f : Type u → Type v} {α : Type u} [Functor f] (x : f α) : f PUnit :=
  Functor.mapConst PUnit.unit x
Discard operation for functors
Informal description
Given a functor `f` and a value `x : f α`, the `discard` operation maps `x` to a value in `f PUnit`, effectively discarding the original value while preserving the functor's structure. This is particularly useful in contexts where only the effect of the functor is desired, not the contained value.
Alternative structure
(f : Type u → Type v) : Type (max (u+1) v) extends Applicative f
Full source
/--
An `Alternative` functor is an `Applicative` functor that can "fail" or be "empty"
and a binary operation `<|>` that “collects values” or finds the “left-most success”.

Important instances include
* `Option`, where `failure := none` and `<|>` returns the left-most `some`.
* Parser combinators typically provide an `Applicative` instance for error-handling and
  backtracking.

Error recovery and state can interact subtly. For example, the implementation of `Alternative` for `OptionT (StateT σ Id)` keeps modifications made to the state while recovering from failure, while `StateT σ (OptionT Id)` discards them.
-/
-- NB: List instance is in mathlib. Once upstreamed, add
-- * `List`, where `failure` is the empty list and `<|>` concatenates.
class Alternative (f : Type u → Type v) : Type (max (u+1) v) extends Applicative f where
  /--
  Produces an empty collection or recoverable failure.  The `<|>` operator collects values or recovers
  from failures. See `Alternative` for more details.
  -/
  failure : {α : Type u} → f α
  /--
  Depending on the `Alternative` instance, collects values or recovers from `failure`s by
  returning the leftmost success. Can be written using the `<|>` operator syntax.
  -/
  orElse  : {α : Type u} → f α → (Unit → f α) → f α
Alternative Functor
Informal description
An `Alternative` functor is an `Applicative` functor equipped with two operations: a constant `failure` representing an empty or failed computation, and a binary operation `<|>` that combines two computations, returning the first successful result. This structure is useful for modeling computations that can fail or be empty, such as optional values or parsers with backtracking.
instOrElseOfAlternative instance
(f : Type u → Type v) (α : Type u) [Alternative f] : OrElse (f α)
Full source
instance (f : Type u → Type v) (α : Type u) [Alternative f] : OrElse (f α) := ⟨Alternative.orElse⟩
Alternative Functor Induces OrElse Operation
Informal description
For any type constructor $f$ that is an `Alternative` functor and any type $\alpha$, there is a canonical `OrElse` structure on $f \alpha$ where the "or else" operation is defined using the `<|>` operation from the `Alternative` instance.
guard definition
{f : Type → Type v} [Alternative f] (p : Prop) [Decidable p] : f Unit
Full source
/--
If the proposition `p` is true, does nothing, else fails (using `failure`).
-/
@[always_inline, inline] def guard {f : Type → Type v} [Alternative f] (p : Prop) [Decidable p] : f Unit :=
  if p then pure () else failure
Guard condition in an Alternative functor
Informal description
Given a proposition $p$ and a type constructor $f$ that is an `Alternative` functor, the function `guard` checks whether $p$ is true (using a `Decidable` instance). If $p$ is true, it returns the unit value `()` in the functor $f$; otherwise, it fails with `failure`.
optional definition
(x : f α) : f (Option α)
Full source
/--
Returns `some x` if `f` succeeds with value `x`, else returns `none`.
-/
@[always_inline, inline] def optional (x : f α) : f (Option α) :=
  somesome <$> x <|> pure none
Optional computation
Informal description
Given a computation `x` of type `f α` in an `Alternative` functor `f`, the operation `optional x` returns a computation that either succeeds with `some` value from `x` if `x` succeeds, or returns `none` if `x` fails.
ToBool structure
(α : Type u)
Full source
class ToBool (α : Type u) where
  toBool : α → Bool
Boolean conversion type class
Informal description
The structure `ToBool` represents a type class that provides a way to convert values of type `α` into a boolean representation. It serves as a generalization of the boolean conversion operation, allowing any type `α` to define its own boolean interpretation.
instToBoolBool instance
: ToBool Bool
Full source
instance : ToBool Bool where
  toBool b := b
Boolean Identity Conversion
Informal description
The Boolean type `Bool` has a canonical instance of the `ToBool` type class, which provides the identity conversion from `Bool` to `Bool`.
bool definition
{β : Type u} {α : Type v} [ToBool β] (f t : α) (b : β) : α
Full source
@[macro_inline] def bool {β : Type u} {α : Type v} [ToBool β] (f t : α) (b : β) : α :=
  match toBool b with
  | true  => t
  | false => f
Boolean conditional selection
Informal description
Given a type `β` with a boolean conversion (via `ToBool`), a type `α`, two values `f t : α`, and a value `b : β`, the function returns `t` if `b` converts to `true`, and `f` otherwise.
orM definition
{m : Type u → Type v} {β : Type u} [Monad m] [ToBool β] (x y : m β) : m β
Full source
/--
Converts the result of the monadic action `x` to a `Bool`. If it is `true`, returns it and ignores
`y`; otherwise, runs `y` and returns its result.

This a monadic counterpart to the short-circuiting `||` operator, usually accessed via the `<||>`
operator.
-/
@[macro_inline] def orM {m : Type u → Type v} {β : Type u} [Monad m] [ToBool β] (x y : m β) : m β := do
  let b ← x
  match toBool b with
  | true  => pure b
  | false => y
Monadic short-circuiting OR
Informal description
Given a monad `m` and a type `β` with a boolean conversion, the function `orM` takes two monadic actions `x` and `y` of type `m β`. It first executes `x` and converts its result to a boolean. If the result is `true`, it returns the result of `x`; otherwise, it executes `y` and returns its result. This function is a monadic version of the short-circuiting logical OR operation.
term_<||>_ definition
: Lean.TrailingParserDescr✝
Full source
infixr:30 " <||> " => orM
Monadic short-circuiting OR operator (`<||>`)
Informal description
The infix operator `<||>` with precedence 30 represents the monadic short-circuiting OR operation, defined as `orM`. Given two monadic actions, it executes the first and returns its result if it evaluates to `true`; otherwise, it executes and returns the second action.
andM definition
{m : Type u → Type v} {β : Type u} [Monad m] [ToBool β] (x y : m β) : m β
Full source
/--
Converts the result of the monadic action `x` to a `Bool`. If it is `true`, returns `y`; otherwise,
returns the original result of `x`.

This a monadic counterpart to the short-circuiting `&&` operator, usually accessed via the `<&&>`
operator.
-/
@[macro_inline] def andM {m : Type u → Type v} {β : Type u} [Monad m] [ToBool β] (x y : m β) : m β := do
  let b ← x
  match toBool b with
  | true  => y
  | false => pure b
Monadic short-circuiting AND
Informal description
Given a monad `m` and a type `β` with a boolean conversion, the function `andM` takes two monadic actions `x : m β` and `y : m β`. It first executes `x` and converts its result to a boolean. If the result is `true`, it executes `y`; otherwise, it returns the original result of `x`. This operation is a monadic version of the short-circuiting logical AND operation.
term_<&&>_ definition
: Lean.TrailingParserDescr✝
Full source
infixr:35 " <&&> " => andM
Monadic AND operator (`<&&>`)
Informal description
The infix operator `<&&>` with precedence 35 is defined as a shorthand for the monadic short-circuiting AND operation `andM`. It takes two monadic actions and performs a logical AND operation where the second action is only executed if the first action evaluates to `true`.
notM definition
{m : Type → Type v} [Applicative m] (x : m Bool) : m Bool
Full source
/--
Runs a monadic action and returns the negation of its result.
-/
@[macro_inline] def notM {m : Type → Type v} [Applicative m] (x : m Bool) : m Bool :=
  notnot <$> x
Monadic boolean negation
Informal description
Given an applicative functor `m` and a monadic action `x : m Bool`, the function `notM` applies the boolean negation to the result of `x`, returning a new monadic action of type `m Bool`.
MonadControl structure
(m : semiOutParam (Type u → Type v)) (n : Type u → Type w)
Full source
/--
A way to lift a computation from one monad to another while providing the lifted computation with a
means of interpreting computations from the outer monad. This provides a means of lifting
higher-order operations automatically.

Clients should typically use `control` or `controlAt`, which request an instance of `MonadControlT`:
the reflexive, transitive closure of `MonadControl`. New instances should be defined for
`MonadControl` itself.
-/
-- This is the same as
-- [`MonadBaseControl`](https://hackage.haskell.org/package/monad-control-1.0.3.1/docs/Control-Monad-Trans-Control.html#t:MonadBaseControl)
-- in Haskell.
class MonadControl (m : semiOutParam (Type u → Type v)) (n : Type u → Type w) where
  /--
  A type that can be used to reconstruct both a returned value and any state used by the outer
  monad.
  -/
  stM      : Type u → Type u
  /--
  Lifts an action from the inner monad `m` to the outer monad `n`. The inner monad has access to a
  reverse lifting operator that can run an `n` action, returning a value and state together.
  -/
  liftWith : {α : Type u} → (({β : Type u} → n β → m (stM β)) → m α) → n α
  /--
  Lifts a monadic action that returns a state and a value in the inner monad to an action in the
  outer monad. The extra state information is used to restore the results of effects from the
  reverse lift passed to `liftWith`'s parameter.
  -/
  restoreM : {α : Type u} → m (stM α) → n α
Monad Control Structure
Informal description
The structure `MonadControl m n` provides a way to lift computations from a base monad `m` to an outer monad `n`, while also allowing the lifted computation to interpret computations from `n` back into `m`. This enables the lifting of higher-order operations automatically, particularly useful when working with monad transformer towers. Key components include: - A function `runInBase : ∀ {α}, n α → m (stM α)` that runs a computation in `n` and packages its result and state in the base monad `m`. - A function `restoreM : m (stM α) → n α` that reconstructs the computation in `n` from the packaged result and state. Here, `stM α` represents the packaged result and state type, which depends on the monad transformer stack.
MonadControlT structure
(m : Type u → Type v) (n : Type u → Type w)
Full source
/--
A way to lift a computation from one monad to another while providing the lifted computation with a
means of interpreting computations from the outer monad. This provides a means of lifting
higher-order operations automatically.

Clients should typically use `control` or `controlAt`, which request an instance of `MonadControlT`:
the reflexive, transitive closure of `MonadControl`. New instances should be defined for
`MonadControl` itself.
-/
class MonadControlT (m : Type u → Type v) (n : Type u → Type w) where
  /--
  A type that can be used to reconstruct both a returned value and any state used by the outer
  monad.
  -/
  stM      : Type u → Type u
  /--
  Lifts an action from the inner monad `m` to the outer monad `n`. The inner monad has access to a
  reverse lifting operator that can run an `n` action, returning a value and state together.
  -/
  liftWith : {α : Type u} → (({β : Type u} → n β → m (stM β)) → m α) → n α
  /--
  Lifts a monadic action that returns a state and a value in the inner monad to an action in the
  outer monad. The extra state information is used to restore the results of effects from the
  reverse lift passed to `liftWith`'s parameter.
  -/
  restoreM {α : Type u} : stM α → n α
Transitive Closure of Monad Control
Informal description
The structure `MonadControlT m n` provides a way to lift computations from a base monad `m` to an outer monad `n`, while also allowing the lifted computation to interpret computations from `n` back into `m`. This enables the lifting of higher-order operations automatically, particularly useful when working with monad transformer towers. Key components include: - A way to run computations in `n` within the base monad `m` (via `runInBase`). - A way to restore the state of the outer monad `n` after running a computation in `m` (via `restoreM`). This structure generalizes the idea of `MonadControl`, forming its reflexive, transitive closure.
instMonadControlTOfMonadControl instance
(m n o) [MonadControl n o] [MonadControlT m n] : MonadControlT m o
Full source
@[always_inline]
instance (m n o) [MonadControl n o] [MonadControlT m n] : MonadControlT m o where
  stM α := stM m n (MonadControl.stM n o α)
  liftWith f := MonadControl.liftWith fun x₂ => liftWith fun x₁ => f (x₁ ∘ x₂)
  restoreM := MonadControl.restoreMMonadControl.restoreM ∘ restoreM
Transitive Composition of Monad Control
Informal description
Given monads $m$, $n$, and $o$, if there exists a `MonadControl` instance from $n$ to $o$ and a `MonadControlT` instance from $m$ to $n$, then there is a `MonadControlT` instance from $m$ to $o$. This means that control operations can be transitively composed across monad transformer stacks.
instMonadControlTOfPure instance
(m : Type u → Type v) [Pure m] : MonadControlT m m
Full source
instance (m : Type u → Type v) [Pure m] : MonadControlT m m where
  stM α := α
  liftWith f := f fun x => x
  restoreM x := pure x
Self-Control in Pure Monads
Informal description
For any monad `m` equipped with a `Pure` instance, there is a canonical `MonadControlT` instance from `m` to itself. This means that computations in `m` can be controlled within the same monad `m`, allowing the lifting and interpretation of computations without changing the monadic context.
controlAt definition
(m : Type u → Type v) {n : Type u → Type w} [MonadControlT m n] [Bind n] {α : Type u} (f : ({β : Type u} → n β → m (stM m n β)) → m (stM m n α)) : n α
Full source
/--
Lifts an operation from an inner monad to an outer monad, providing it with a reverse lifting
operator that allows outer monad computations to be run in the inner monad. The lifted operation is
required to return extra information that is required in order to reconstruct the reverse lift's
effects in the outer monad; this extra information is determined by `stM`.

This function takes the inner monad as an explicit parameter. Use `control` to infer the monad.
-/
@[always_inline, inline]
def controlAt (m : Type u → Type v) {n : Type u → Type w} [MonadControlT m n] [Bind n] {α : Type u}
    (f : ({β : Type u} → n β → m (stM m n β)) → m (stM m n α)) : n α :=
  liftWithliftWith f >>= restoreM
Monadic control with explicit base monad
Informal description
Given a base monad `m` and an outer monad `n` with `MonadControlT m n` and `Bind n` instances, the function `controlAt` lifts an operation from `m` to `n`. It takes a function `f` that maps computations in `n` to computations in `m` (via `runInBase`) and produces a result in `m` that includes the necessary state information `stM m n α` to reconstruct the effect in `n`. The result is then restored to the outer monad `n` using `restoreM`. More precisely, for any type `α`, the function `f` takes a polymorphic function `runInBase : ∀ β, n β → m (stM m n β)` and returns a computation in `m` producing a packaged result `stM m n α`. The `controlAt` function then lifts this to `n α` by first applying `f` and then restoring the monadic context using `restoreM`.
control definition
{m : Type u → Type v} {n : Type u → Type w} [MonadControlT m n] [Bind n] {α : Type u} (f : ({β : Type u} → n β → m (stM m n β)) → m (stM m n α)) : n α
Full source
/--
Lifts an operation from an inner monad to an outer monad, providing it with a reverse lifting
operator that allows outer monad computations to be run in the inner monad. The lifted operation is
required to return extra information that is required in order to reconstruct the reverse lift's
effects in the outer monad; this extra information is determined by `stM`.

This function takes the inner monad as an implicit parameter. Use `controlAt` to specify it
explicitly.
-/
@[always_inline, inline]
def control {m : Type u → Type v} {n : Type u → Type w} [MonadControlT m n] [Bind n] {α : Type u}
    (f : ({β : Type u} → n β → m (stM m n β)) → m (stM m n α)) : n α :=
  controlAt m f
Monadic control operation
Informal description
Given a base monad `m` and an outer monad `n` with `MonadControlT m n` and `Bind n` instances, the function `control` lifts an operation from `m` to `n`. It takes a function `f` that maps computations in `n` to computations in `m` (via `runInBase`) and produces a result in `m` that includes the necessary state information `stM m n α` to reconstruct the effect in `n`. The result is then restored to the outer monad `n` using `restoreM`. More precisely, for any type `α`, the function `f` takes a polymorphic function `runInBase : ∀ β, n β → m (stM m n β)` and returns a computation in `m` producing a packaged result `stM m n α`. The `control` function then lifts this to `n α` by first applying `f` and then restoring the monadic context using `restoreM`.
ForM structure
(m : Type u → Type v) (γ : Type w₁) (α : outParam (Type w₂))
Full source
/--
Overloaded monadic iteration over some container type.

An instance of `ForM m γ α` describes how to iterate a monadic operator over a container of type `γ`
with elements of type `α` in the monad `m`. The element type should be uniquely determined by the
monad and the container.

Use `ForM.forIn` to construct a `ForIn` instance from a `ForM` instance, thus enabling the use of
the `for` operator in `do`-notation.
-/
class ForM (m : Type u → Type v) (γ : Type w₁) (α : outParam (Type w₂)) where
  /--
  Runs the monadic action `f` on each element of the collection `coll`.
  -/
  forM [Monad m] (coll : γ) (f : α → m PUnit) : m PUnit
Monadic iteration over a container
Informal description
The structure `ForM m γ α` describes how to iterate a monadic operation over a container of type `γ` with elements of type `α` in the monad `m`. The element type $\alpha$ is uniquely determined by the monad $m$ and the container type $\gamma$. This enables monadic iteration patterns over container types.
Bind.kleisliRight definition
[Bind m] (f₁ : α → m β) (f₂ : β → m γ) (a : α) : m γ
Full source
/-- Left-to-right composition of Kleisli arrows. -/
@[always_inline]
def Bind.kleisliRight [Bind m] (f₁ : α → m β) (f₂ : β → m γ) (a : α) : m γ :=
  f₁ a >>= f₂
Left-to-right Kleisli composition
Informal description
Given a monad `m` and two Kleisli arrows `f₁ : α → m β` and `f₂ : β → m γ`, the function `Bind.kleisliRight` composes them left-to-right, producing a new Kleisli arrow `α → m γ`. Specifically, for any input `a : α`, it first applies `f₁` to `a` and then binds the result with `f₂`. This is equivalent to the composition `f₁ >=> f₂` in the Kleisli category of the monad `m`.
Bind.kleisliLeft definition
[Bind m] (f₂ : β → m γ) (f₁ : α → m β) (a : α) : m γ
Full source
/-- Right-to-left composition of Kleisli arrows. -/
@[always_inline]
def Bind.kleisliLeft [Bind m] (f₂ : β → m γ) (f₁ : α → m β) (a : α) : m γ :=
  f₁ a >>= f₂
Right-to-left Kleisli composition
Informal description
Given a monad `m` and two Kleisli arrows `f₁ : α → m β` and `f₂ : β → m γ`, the function `Bind.kleisliLeft` composes them right-to-left, producing a new Kleisli arrow `α → m γ`. Specifically, for any input `a : α`, it first applies `f₁` to `a` and then binds the result with `f₂`. This is equivalent to the composition `f₂ ∘ f₁` in the Kleisli category of the monad `m`.
Bind.bindLeft definition
[Bind m] (f : α → m β) (ma : m α) : m β
Full source
/-- Same as `Bind.bind` but with arguments swapped. -/
@[always_inline]
def Bind.bindLeft [Bind m] (f : α → m β) (ma : m α) : m β :=
  ma >>= f
Left bind operation in a monad
Informal description
Given a monad `m`, a function `f : α → m β`, and a monadic value `ma : m α`, the function `Bind.bindLeft` applies `f` to the result of `ma` using the bind operation `>>=`. This is equivalent to `ma >>= f`.
term_>=>_ definition
: Lean.TrailingParserDescr✝
Full source
@[inherit_doc] infixr:55 " >=> " => Bind.kleisliRight
Right-to-left Kleisli composition operator (`>=>`)
Informal description
The infix operator `>=>` represents right-to-left Kleisli composition in a monad. Given two Kleisli arrows `f : α → m β` and `g : β → m γ`, the composition `f >=> g` yields a new Kleisli arrow `α → m γ` that first applies `f` and then `g` to the result.
term_<=<_ definition
: Lean.TrailingParserDescr✝
Full source
@[inherit_doc] infixr:55 " <=< " => Bind.kleisliLeft
Kleisli left composition operator
Informal description
The infix operator `<=<` (Kleisli left composition) with precedence 55 is defined as right-associative composition in the Kleisli category, equivalent to `Bind.kleisliLeft`. For monads, this represents composition of monadic functions in reverse order: `(f <=< g) x` is equivalent to `g x >>= f`.
term_=<<_ definition
: Lean.TrailingParserDescr✝
Full source
@[inherit_doc] infixr:55 " =<< " => Bind.bindLeft
Monadic bind operator `=<<`
Informal description
The infix operator `=<<` with precedence 55 is defined as right-associative and stands for the monadic bind operation `Bind.bindLeft`. It is used to chain monadic computations where the function is applied to the value inside the monad and returns a new monadic computation.