Idris2Doc : Language.IntrinsicTyping.STLCR

Language.IntrinsicTyping.STLCR(source)

The content of this module is based on the MSc Thesis
Coinductive Formalization of SECD Machine in Agda
by Adam Krupička

Definitions

dataSTLCR : List (Ty, Ty) ->ListTy->Ty->Type
Totality: total
Visibility: public export
Constructors:
Var : Elemtyg->STLCRrgty
Lam : STLCR ((a, b) ::r) (a::g) b->STLCRrg (TyFunab)
App : STLCRrg (TyFunab) ->STLCRrga->STLCRrgb
Rec : Elem (a, b) r->STLCRrg (TyFunab)
If : STLCRrgTyBool->STLCRrga->STLCRrga->STLCRrga
Eqb : STLCRrga->STLCRrga->STLCRrgTyBool
Lit : Constty->STLCRrgty
Add : STLCRrgTyInt->STLCRrgTyInt->STLCRrgTyInt
Sub : STLCRrgTyInt->STLCRrgTyInt->STLCRrgTyInt
Mul : STLCRrgTyInt->STLCRrgTyInt->STLCRrgTyInt
fromInteger : Integer->STLCRrgTyInt
Totality: total
Visibility: public export
compile : STLCRrgty->Steps (MkStatesgr) (MkState (ty::s) gr)
Totality: total
Visibility: public export
compileT : STLCR ((a, b) ::r) (a::g) b->Steps (MkState [] (a::g) ((a, b) ::r)) (MkState [b] (a::g) ((a, b) ::r))
Totality: total
Visibility: public export
compileAcc : Stepzinit (MkStatesgr) ->STLCRrgty->Stepzinit (MkState (ty::s) gr)
Totality: total
Visibility: public export