Monadic lazy lists. #
Lazy lists with "laziness" controlled by an arbitrary monad.
In an initial section we describe the specification of MLList
,
and provide a private unsafe implementation,
and then a public opaque
wrapper of this implementation, satisfying the specification.
Constructs a MLList
from head and tail.
Equations
- MLList.cons = (MLList.spec✝ m).cons
Instances For
Embed a non-monadic thunk as a lazy list.
Equations
Instances For
Lift a monadic lazy list inside the monad to a monadic lazy list.
Equations
Instances For
Equations
- MLList.instEmptyCollection = { emptyCollection := MLList.nil }
Equations
- MLList.instInhabited = { default := MLList.nil }
Equations
- MLList.instForInOfMonadOfMonadLiftT = { forIn := fun {β : Type ?u.40} [Monad n] => MLList.forIn }
Construct a singleton monadic lazy list from a single monadic value.
Equations
- MLList.singletonM x = MLList.squash fun (x_1 : Unit) => do let __do_lift ← x pure (MLList.cons __do_lift MLList.nil)
Instances For
Construct a singleton monadic lazy list from a single value.
Equations
Instances For
Constructs an MLList
recursively, with state in α
, recording terms from β
.
If f
returns none
the list will terminate.
Variant of MLList.fix?
that allows returning values of a different type.
Repeatedly apply a function f : α → m (α × List β)
to an initial a : α
,
accumulating the elements of the resulting List β
as a single monadic lazy list.
(This variant allows starting with a specified List β
of elements, as well. )
Repeatedly apply a function f : α → m (α × List β)
to an initial a : α
,
accumulating the elements of the resulting List β
as a single monadic lazy list.
Equations
- MLList.fixl f s = MLList.fixlWith f s []
Instances For
Equations
- MLList.ofList [] = MLList.nil
- MLList.ofList (h :: t) = MLList.cons h (MLList.thunk fun (x : Unit) => MLList.ofList t)
Instances For
Convert a List
of values inside the monad into a MLList
.
Equations
- MLList.ofListM [] = MLList.nil
- MLList.ofListM (h :: t) = MLList.squash fun (x : Unit) => do let __do_lift ← h pure (MLList.cons __do_lift (MLList.ofListM t))
Instances For
Performs a case distinction on a MLList
when the motive is a MLList
as well.
(We need to be in a monadic context to distinguish a nil from a cons.)
Equations
- One or more equations did not get rendered due to their size.
Instances For
Gives the monadic lazy list consisting all of folds of a function on a given initial element.
Thus [a₀, a₁, ...].foldsM f b
will give [b, ← f b a₀, ← f (← f b a₀) a₁, ...]
.
Gives the monadic lazy list consisting all of folds of a function on a given initial element.
Thus [a₀, a₁, ...].foldsM f b
will give [b, f b a₀, f (f b a₀) a₁, ...]
.
Equations
- MLList.folds f init L = MLList.foldsM (fun (b : β) (a : α) => pure (f b a)) init L
Instances For
Take the first n
elements, as a list inside the monad.
Equations
- xs.takeAsList n = MLList.takeAsList.go n [] xs
Instances For
Take the first n
elements, as an array inside the monad.
Equations
- xs.takeAsArray n = MLList.takeAsArray.go n #[] xs
Instances For
Apply a function to every element of a MLList
.
Equations
- MLList.map f L = MLList.mapM (fun (a : α) => pure (f a)) L
Instances For
Filter a MLList
.
Equations
- MLList.filter p L = MLList.filterM (fun (a : α) => pure { down := p a }) L
Instances For
Filter and transform a MLList
using an Option
valued function.
Equations
- MLList.filterMap f = MLList.filterMapM fun (a : α) => pure (f a)
Instances For
Take the initial segment of the lazy list, until the function f
first returns false
.
Equations
- MLList.takeWhile f = MLList.takeWhileM fun (a : α) => pure { down := f a }
Instances For
Enumerate the elements of a monadic lazy list.
Equations
Instances For
The infinite monadic lazy list of natural numbers.
Equations
- MLList.range = MLList.fix (fun (n : Nat) => pure (n + 1)) 0
Instances For
Implementation of MLList.fin
.
Convert an array to a monadic lazy list.
Equations
- MLList.ofArray L = MLList.ofArray.go L 0
Instances For
Implementation of MLList.ofArray
.
Group the elements of a lazy list into chunks of a given size. If the lazy list is finite, the last chunk may be smaller (possibly even length 0).
Equations
- L.chunk n = MLList.chunk.go n n #[] L
Instances For
Add one element to the end of a monadic lazy list.
Equations
- L.concat a = L.append fun (x : Unit) => MLList.cons a MLList.nil
Instances For
Convert any value in the monad to the singleton monadic lazy list.
Equations
- MLList.monadLift x = MLList.squash fun (x_1 : Unit) => do let __do_lift ← x pure (MLList.cons __do_lift MLList.nil)
Instances For
Lift the monad of a lazy list.
Run a lazy list in a StateRefT'
monad on some initial state.
Return the head of a monadic lazy list if it exists, as an Option
in the monad.
Equations
Instances For
Take the initial segment of the lazy list,
up to and including the first place where f
gives true
.
Equations
- L.takeUpToFirst f = L.takeUpToFirstM fun (a : α) => pure { down := f a }
Instances For
Gets the last element of a monadic lazy list, or the default value if the list is empty. This will run forever if the list is infinite.
Equations
- L.getLast! = Option.get! <$> L.getLast?
Instances For
Folds a binary function across a monadic lazy list, from an initial starting value. This will run forever if the list is infinite.
Equations
- MLList.foldM f init L = do let __do_lift ← (MLList.foldsM f init L).getLast? pure (__do_lift.getD init)
Instances For
Folds a binary function across a monadic lazy list, from an initial starting value. This will run forever if the list is infinite.
Equations
- MLList.fold f init L = MLList.foldM (fun (b : β) (a : α) => pure (f b a)) init L
Instances For
Return the head of a monadic lazy list, as a value in the monad. Fails if the list is empty.
Equations
Instances For
Apply a function returning values inside the monad to a monadic lazy list, returning only the first successful result.
Equations
- L.firstM f = (MLList.filterMapM f L).head
Instances For
Return the first value on which a predicate returns true.
Equations
- L.first p = (MLList.filter p L).head
Instances For
Equations
Equations
- MLList.instAlternativeOfMonad = Alternative.mk (fun {α : Type ?u.21} => MLList.nil) fun {α : Type ?u.21} => MLList.append
Equations
- MLList.instMonadLiftOfMonad = { monadLift := fun {α : Type ?u.20} => MLList.monadLift }