cek-call-cc/src/Project/Definitions.agda
2021-12-09 06:02:52 -06:00

107 lines
No EOL
2.7 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Project.Definitions where
open import Data.Maybe using (Maybe; just; nothing)
open import Data.Nat using (; zero; suc)
open import Data.Product renaming (_,_ to _ʻ_)
open import Project.Util using (_$_)
infix 4 _∋_
infix 4 _·_
infix 4 _∘_
infixl 6 _,_
infixr 7 _⇒_
data Type : Set where
: Type
` : Type
_⇒_ : Type Type Type
K[_⇒_] : Type Type Type
data Context : Set where
: Context
_,_ : Context Type Context
data Value ( : Type) : Type Set
data Env ( : Type) : Context Set where
: Env
_[__] : {Γ} Env Γ (A : Type) Value A Env (Γ , A)
data _∋_ : Context Type Set where
zero : {Γ A} Γ , A A
suc : {Γ A B} Γ A Γ , B A
lookup : { Γ A} Env Γ Γ A Value A
lookup ()
lookup (env [ A x ]) zero = x
lookup (env [ A x ]) (suc id) = lookup env id
data Aexp ( : Type) Context : Type Set
data Exp ( : Type) Context : Type Set
data Aexp Γ where
value : {A} Value A Aexp Γ A
-- Natural numbers
zero : Aexp Γ `
suc : Aexp Γ ` Aexp Γ `
-- Functions
`_ : {A} Γ A Aexp Γ A
ƛ : {B} {A : Type} Exp (Γ , A) B Aexp Γ (A B)
data Exp Γ where
abort : {A} Aexp Γ Exp Γ A
-- Atomic expressions
atomic : {A} Aexp Γ A Exp Γ A
-- Natural numbers
case : {A} Aexp Γ ` Exp Γ A Aexp Γ (` A) Exp Γ A
-- Functions
_·_ : {A B} Aexp Γ (A B) Aexp Γ A Exp Γ B
_∘_ : {A} Aexp Γ K[ A ] Aexp Γ A Exp Γ
-- Call/cc
call/cc : {A} Aexp Γ (K[ A ] ) Exp Γ A
-- Let
`let : {A B : Type} Exp Γ A Exp (Γ , A) B Exp Γ B
data Kont ( : Type) : Type Set
record Letk (Tv : Type) : Set
data Value where
-- Natural numbers
zero : Value `
suc : Value ` Value `
-- Functions
clo : {Γ} {A B : Type} Exp (Γ , A) B Env Γ Value (A B)
-- Call/CC
cont : {A} Kont A Value K[ A ]
record Letk Tv where
inductive
constructor letk
field
{Tc} : Type
Γ : Context
C : Exp (Γ , Tv) Tc
E : Env Γ
K : Kont Tc
data Kont where
halt : Kont
kont : {Tc} Letk Tc Kont Tc
record State ( : Type) : Set where
constructor mkState
field
Tc : Type
Γ : Context
C : Exp Γ Tc
E : Env Γ
K : Kont Tc