cek-call-cc/src/Project/Definitions.agda

107 lines
2.7 KiB
Agda
Raw Normal View History

2021-12-08 06:33:28 +00:00
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 _·_
2021-12-09 11:28:25 +00:00
infix 4 _∘_
2021-12-08 06:33:28 +00:00
infixl 6 _,_
infixr 7 _⇒_
data Type : Set where
2021-12-09 08:05:20 +00:00
: Type
2021-12-08 06:33:28 +00:00
` : Type
2021-12-09 10:32:32 +00:00
_⇒_ : Type Type Type
2021-12-09 11:28:25 +00:00
K[_⇒_] : Type Type Type
2021-12-08 06:33:28 +00:00
data Context : Set where
: Context
_,_ : Context Type Context
2021-12-09 12:00:09 +00:00
data Value ( : Type) : Type Set
2021-12-08 06:33:28 +00:00
2021-12-09 12:00:09 +00:00
data Env ( : Type) : Context Set where
: Env
_[__] : {Γ} Env Γ (A : Type) Value A Env (Γ , A)
2021-12-08 06:33:28 +00:00
data _∋_ : Context Type Set where
zero : {Γ A} Γ , A A
suc : {Γ A B} Γ A Γ , B A
2021-12-09 12:00:09 +00:00
lookup : { Γ A} Env Γ Γ A Value A
2021-12-08 06:33:28 +00:00
lookup ()
lookup (env [ A x ]) zero = x
lookup (env [ A x ]) (suc id) = lookup env id
2021-12-09 12:00:09 +00:00
data Aexp ( : Type) Context : Type Set
data Exp ( : Type) Context : Type Set
2021-12-08 06:33:28 +00:00
2021-12-09 12:00:09 +00:00
data Aexp Γ where
value : {A} Value A Aexp Γ A
2021-12-09 08:05:20 +00:00
2021-12-08 06:33:28 +00:00
-- Natural numbers
2021-12-09 12:00:09 +00:00
zero : Aexp Γ `
suc : Aexp Γ ` Aexp Γ `
2021-12-08 06:33:28 +00:00
-- Functions
2021-12-09 12:00:09 +00:00
`_ : {A} Γ A Aexp Γ A
ƛ : {B} {A : Type} Exp (Γ , A) B Aexp Γ (A B)
2021-12-08 06:33:28 +00:00
2021-12-09 12:00:09 +00:00
data Exp Γ where
abort : {A} Aexp Γ Exp Γ A
2021-12-09 10:32:32 +00:00
2021-12-08 06:33:28 +00:00
-- Atomic expressions
2021-12-09 12:00:09 +00:00
atomic : {A} Aexp Γ A Exp Γ A
2021-12-08 06:33:28 +00:00
-- Natural numbers
2021-12-09 12:00:09 +00:00
case : {A} Aexp Γ ` Exp Γ A Aexp Γ (` A) Exp Γ A
2021-12-08 06:33:28 +00:00
-- Functions
2021-12-09 12:00:09 +00:00
_·_ : {A B} Aexp Γ (A B) Aexp Γ A Exp Γ B
_∘_ : {A} Aexp Γ K[ A ] Aexp Γ A Exp Γ
2021-12-08 06:33:28 +00:00
-- Call/cc
2021-12-09 12:00:09 +00:00
call/cc : {A} Aexp Γ (K[ A ] ) Exp Γ A
2021-12-09 10:32:32 +00:00
-- Let
2021-12-09 12:00:09 +00:00
`let : {A B : Type} Exp Γ A Exp (Γ , A) B Exp Γ B
2021-12-09 10:32:32 +00:00
2021-12-08 06:33:28 +00:00
data Kont ( : Type) : Type Set
2021-12-09 08:05:20 +00:00
record Letk (Tv : Type) : Set
2021-12-08 06:33:28 +00:00
2021-12-09 12:00:09 +00:00
data Value where
2021-12-08 06:33:28 +00:00
-- Natural numbers
2021-12-09 12:00:09 +00:00
zero : Value `
suc : Value ` Value `
2021-12-08 06:33:28 +00:00
-- Functions
2021-12-09 12:00:09 +00:00
clo : {Γ} {A B : Type} Exp (Γ , A) B Env Γ Value (A B)
2021-12-08 06:33:28 +00:00
-- Call/CC
2021-12-09 12:00:09 +00:00
cont : {A} Kont A Value K[ A ]
2021-12-08 06:33:28 +00:00
2021-12-09 08:05:20 +00:00
record Letk Tv where
2021-12-08 06:33:28 +00:00
inductive
constructor letk
field
{Tc} : Type
Γ : Context
2021-12-09 12:00:09 +00:00
C : Exp (Γ , Tv) Tc
E : Env Γ
2021-12-08 06:33:28 +00:00
K : Kont Tc
data Kont where
halt : Kont
2021-12-09 08:05:20 +00:00
kont : {Tc} Letk Tc Kont Tc
2021-12-08 06:33:28 +00:00
record State ( : Type) : Set where
constructor mkState
field
Tc : Type
Γ : Context
2021-12-09 12:00:09 +00:00
C : Exp Γ Tc
E : Env Γ
2021-12-08 06:33:28 +00:00
K : Kont Tc