Idris2Doc : Data.Recursion.Free

Data.Recursion.Free(source)

Module partially based on McBride's paper:
Turing-Completeness Totally Free

It gives us a type to describe computation using general recursion
and functions to run these computations for a while or to completion
if we are able to prove them total.

The content of the Erased section is new. Instead of producing the
domain/evaluation pair by computing a Dybjer-Setzer code we build a
specialised structure that allows us to make the domain proof runtime
irrelevant.

Definitions

dataGeneral : (a : Type) -> (a->Type) ->Type->Type
  Syntax for a program using general recursion

Totality: total
Visibility: public export
Constructors:
Tell : x->Generalabx
  We can return a value without performing any recursive call.
Ask : (i : a) -> (bi->Generalabx) ->Generalabx
  Or we can pick an input and ask an oracle to give us a return value
for it. The second argument is a continuation explaining what we want
to do with the returned value.

Hints:
Applicative (Generalab)
Functor (Generalab)
Monad (Generalab)
PiG : (a : Type) -> (a->Type) ->Type
  Type of functions using general recursion

Totality: total
Visibility: public export
fold : (x->y) -> ((i : a) -> (bi->y) ->y) ->Generalabx->y
  Recursor for General

Totality: total
Visibility: public export
call : PiGab
  Perform a recursive call and return the value provided by the oracle.

Totality: total
Visibility: public export
bind : Generalabx-> (x->Generalaby) ->Generalaby
  Monadic bind (defined outside of the interface to be able to use it for
map and (<*>)).

Totality: total
Visibility: public export
monadMorphism : Monadm=> ((i : a) ->m (bi)) ->Generalabx->mx
  Given a monadic oracle we can give a monad morphism interpreting a
function using general recursion as a monadic process.

Totality: total
Visibility: public export
already : Generalabx->Maybex
  Check whether we are ready to return a value

Totality: total
Visibility: public export
expand : PiGab->Generalabx->Generalabx
  Use a function using general recursion to expand all of the oracle calls.

Totality: total
Visibility: public export
engine : PiGab->Nat->Generalabx->Generalabx
  Recursively call expand a set number of times.

Totality: total
Visibility: public export
petrol : PiGab->Nat-> (i : a) ->Maybe (bi)
  Check whether recursively calling expand a set number of times is enough
to produce a value.

Totality: total
Visibility: public export
late : PiGab->Generalabx->Latex
  Rely on an oracle using general recursion to convert a function using
general recursion into a process returning a value in the (distant) future.

Totality: total
Visibility: public export
lazy : PiGab-> (i : a) ->Late (bi)
  Interpret a function using general recursion as a process returning
a value in the (distant) future.

Totality: total
Visibility: public export
Domain : PiGab-> (i : a) ->Codeb (bi)
  Compute, as a Dybjer-Setzer code for an inductive-recursive type, the domain
of a function defined by general recursion.

Totality: total
Visibility: public export
evaluate : (f : PiGab) -> (i : a) ->Mu (Domainf) i->bi
  If a given input is in the domain of the function then we may evaluate
it fully on that input and obtain a pure return value.

Totality: total
Visibility: public export
totally : (f : PiGab) -> ((i : a) ->Mu (Domainf) i) -> (i : a) ->bi
  If every input value is in the domain then the function is total.

Totality: total
Visibility: public export
dataLayer : PiGab->Generalab (bi) ->Type
Totality: total
Visibility: public export
Constructors:
MkTell : (o : bi) ->Layerf (Tello)
  A computation returning a value is trivially terminating
MkAsk : (j : a) -> (jprf : Domainfj) -> (k : (bj->Generalab (bi))) ->Layerf (k (evaluatefjjprf)) ->Layerf (Askjk)
  Performing a call to the oracle is termnating if the input is in its
domain and if the rest of the computation is also finite.
view : (d : Generalab (bi)) -> (0l : Layerfd) ->Viewl
  Function computing the view by pattern-matching on the computation and
inverting the proof. Note that the proof is runtime irrelevant even though
the resulting view is not: this is possible because the relevant constructor
is uniquely determined by the shape of `d`.

Totality: total
Visibility: public export
totally : (f : PiGab) -> (0_ : ((i : a) ->Domainfi)) -> (i : a) ->bi
  If a function's domain is total then it is a pure function.

Totality: total
Visibility: public export
irrelevantDomain : (f : PiGab) -> (i : a) -> (p : Domainfi) -> (q : Domainfi) ->p=q
  Domain is a singleton type

Totality: total
Visibility: export
evaluateLayerIrrelevance : (f : PiGab) -> (d : Generalab (bi)) -> (0p : Layerfd) -> (0q : Layerfd) ->evaluateLayerfdp=evaluateLayerfdq
  The result of `evaluateLayer` does not depend on the specific proof that
`i` is in the domain of the layer of computation at hand.

Totality: total
Visibility: export
evaluateIrrelevance : (f : PiGab) -> (i : a) -> (0p : Domainfi) -> (0q : Domainfi) ->evaluatefip=evaluatefiq
  The result of `evaluate` does not depend on the specific proof that `i`
is in the domain of the function at hand.

Totality: total
Visibility: export
totallyIrrelevance : (f : PiGab) -> (0p : ((i : a) ->Domainfi)) -> (0q : ((i : a) ->Domainfi)) -> (i : a) ->totallyfpi=totallyfqi
  The result computed by a total function is independent from the proof
that it is total.

Totality: total
Visibility: export