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 (Tω : Type) : Type → Set data Env (Tω : Type) : Context → Set where ∅ : Env Tω ∅ _[_∶_] : ∀ {Γ} → Env Tω Γ → (A : Type) → Value Tω A → Env Tω (Γ , A) data _∋_ : Context → Type → Set where zero : ∀ {Γ A} → Γ , A ∋ A suc : ∀ {Γ A B} → Γ ∋ A → Γ , B ∋ A lookup : ∀ {Tω Γ A} → Env Tω Γ → Γ ∋ A → Value Tω A lookup ∅ () lookup (env [ A ∶ x ]) zero = x lookup (env [ A ∶ x ]) (suc id) = lookup env id data Aexp (Tω : Type) Context : Type → Set data Exp (Tω : Type) Context : Type → Set data Aexp Tω Γ where value : ∀ {A} → Value Tω A → Aexp Tω Γ A -- Natural numbers zero : Aexp Tω Γ `ℕ suc : Aexp Tω Γ `ℕ → Aexp Tω Γ `ℕ -- Functions `_ : ∀ {A} → Γ ∋ A → Aexp Tω Γ A ƛ : ∀ {B} {A : Type} → Exp Tω (Γ , A) B → Aexp Tω Γ (A ⇒ B) data Exp Tω Γ where abort : ∀ {A} → Aexp Tω Γ ⊥ → Exp Tω Γ A -- Atomic expressions atomic : ∀ {A} → Aexp Tω Γ A → Exp Tω Γ A -- Natural numbers case : ∀ {A} → Aexp Tω Γ `ℕ → Exp Tω Γ A → Aexp Tω Γ (`ℕ ⇒ A) → Exp Tω Γ A -- Functions _·_ : ∀ {A B} → Aexp Tω Γ (A ⇒ B) → Aexp Tω Γ A → Exp Tω Γ B _∘_ : ∀ {A} → Aexp Tω Γ K[ A ⇒ ⊥ ] → Aexp Tω Γ A → Exp Tω Γ ⊥ -- Call/cc call/cc : ∀ {A} → Aexp Tω Γ (K[ A ⇒ ⊥ ] ⇒ Tω) → Exp Tω Γ A -- Let `let : ∀ {A B : Type} → Exp Tω Γ A → Exp Tω (Γ , A) B → Exp Tω Γ B data Kont (Tω : Type) : Type → Set record Letk (Tv Tω : Type) : Set data Value Tω where -- Natural numbers zero : Value Tω `ℕ suc : Value Tω `ℕ → Value Tω `ℕ -- Functions clo : ∀ {Γ} {A B : Type} → Exp Tω (Γ , A) B → Env Tω Γ → Value Tω (A ⇒ B) -- Call/CC cont : ∀ {A} → Kont Tω A → Value Tω K[ A ⇒ ⊥ ] record Letk Tω Tv where inductive constructor letk field {Tc} : Type Γ : Context C : Exp Tω (Γ , Tv) Tc E : Env Tω Γ K : Kont Tω Tc data Kont Tω where halt : Kont Tω Tω kont : ∀ {Tc} → Letk Tω Tc → Kont Tω Tc record State (Tω : Type) : Set where constructor mkState field Tc : Type Γ : Context C : Exp Tω Γ Tc E : Env Tω Γ K : Kont Tω Tc