Idris2Doc : Data.Description.Regular

Data.Description.Regular(source)

The content of this module is based on the paper
A Completely Unique Account of Enumeration
by Cas van der Rest, and Wouter Swierstra
https://doi.org/10.1145/3547636

Definitions

dataDesc : (Type->Type) ->Type
  Description of regular functors
@ p stores additional data for constant types
e.g. the fact they are enumerable

Totality: total
Visibility: public export
Constructors:
Zero : Descp
  The code for the empty type
One : Descp
  The code for the unit type
Id : Descp
  The code for the identity functor
Const : (0s : Type) ->ps->Descp
  The code for the constant functor
(*) : Descp->Descp->Descp
  The code for the product of functors
(+) : Descp->Descp->Descp
  The code for the sum of functors
0Elem : Descp->Type->Type
  Computing the meaning of a description as a functor

Totality: total
Visibility: public export
map : (d : Descp) -> (a->b) ->Elemda->Elemdb
  Elem does decode to functors

Totality: total
Visibility: export
dataFix : Descp->Type
  A regular type is obtained by taking the fixpoint of
the decoding of a description.
Unfortunately Idris 2 does not currently detect that this definition
is total because we do not track positivity in function arguments

Totality: total
Visibility: public export
Constructor: 
MkFix : assert_total (Elemd (Fixd)) ->Fixd
fold : (Elemdx->x) ->Fixd->x
  Fix is an initial algebra so we get the fold

Totality: total
Visibility: export
record(~>) : (d : Descp) -> (Fixd->Type) ->Type
  A memo trie is a coinductive structure

Totality: total
Visibility: export
Constructor: 
MkMemo : assert_total (Memod (\x=> Inf (d~>x)) (b.MkFix)) ->d~>b

Projection: 
.getMemo : d~>b->assert_total (Memod (\x=> Inf (d~>x)) (b.MkFix))

Fixity Declaration: infixr operator, level 0
trie : ((x : Fixd) ->bx) ->d~>b
Totality: total
Visibility: export
untrie : d~>b-> (x : Fixd) ->bx
Totality: total
Visibility: export
memo : (0b : (Fixd->Type)) -> ((x : Fixd) ->bx) -> (x : Fixd) ->bx
Totality: total
Visibility: export