Idris2Doc : Data.ProofDelay

Data.ProofDelay(source)

The contents of this module are based on the paper
Deferring the Details and Deriving Programs
by Liam O'Connor
https://doi.org/10.1145/3331554.3342605

Definitions

recordPDelay : Type->Type
  A type `x` which can only be computed once some, delayed, proof obligations
have been fulfilled.

Totality: total
Visibility: public export
Constructor: 
Prf : (goals : ListType) -> (HListgoals->x) ->PDelayx

Projections:
.goals : PDelayx->ListType
  List of propositions we need to prove.
.prove : ({rec:0} : PDelayx) ->HList (goals{rec:0}) ->x
  Given the proofs required (i.e. the goals), actually compute the value x.
.goals : PDelayx->ListType
  List of propositions we need to prove.

Totality: total
Visibility: public export
goals : PDelayx->ListType
  List of propositions we need to prove.

Totality: total
Visibility: public export
.prove : ({rec:0} : PDelayx) ->HList (goals{rec:0}) ->x
  Given the proofs required (i.e. the goals), actually compute the value x.

Totality: total
Visibility: public export
prove : ({rec:0} : PDelayx) ->HList (goals{rec:0}) ->x
  Given the proofs required (i.e. the goals), actually compute the value x.

Totality: total
Visibility: public export
pure : tx->PDelaytx
  A value which can immediately be constructed (i.e. no delayed proofs).

Totality: total
Visibility: public export
later : PDelaytx
  Delay the full computation of `x` until `later`.

Totality: total
Visibility: public export
(<*>) : PDelay (a->b) ->PDelaya->PDelayb
  We can compose `PDelay` computations.

Totality: total
Visibility: public export
Fixity Declaration: infixl operator, level 3
dataOList : Nat->Nat->Type
  An ordered list type.

[27](https://dl.acm.org/doi/10.1145/2503778.2503786)

Totality: total
Visibility: public export
Constructors:
Nil : LTEmn->OListmn
Cons : (x : Nat) ->LTEmx->OListxn->OListmn
dataBST : Nat->Nat->Type
  A binary search tree carrying proofs of the ordering in the leaves.

[31](https://dl.acm.org/doi/10.1145/2628136.2628163)

Totality: total
Visibility: public export
Constructors:
Leaf : LTEmn->BSTmn
Branch : (x : Nat) ->BSTmx->BSTxn->BSTmn
nil : PDelay (OListmn)
  OList `Nil`, but delaying the proof obligation.

Totality: total
Visibility: public export
cons : (x : Nat) ->PDelay (OListxn) ->PDelay (OListmn)
  OList `Cons`, but delaying the proof obligations.

Totality: total
Visibility: public export
example : OList15
  Extracting an actual `OList` from the delayed version requires providing the
unergonomic proofs.

Totality: total
Visibility: public export
leaf : PDelay (BSTmn)
  BST `Leaf`, but delaying the proof obligation.

Totality: total
Visibility: public export
branch : (x : Nat) ->PDelay (BSTmx) ->PDelay (BSTxn) ->PDelay (BSTmn)
  BST `Branch`, but delaying the proof obligations.

Totality: total
Visibility: public export
example2 : BST210
  Once again, extracting the actual `BST` requires providing the uninteresting
proofs.

Totality: total
Visibility: public export