2024-06-07 14:22:38 +00:00
|
|
|
|
{-# OPTIONS --prop #-}
|
2024-06-06 14:21:04 +00:00
|
|
|
|
module Ahmed.Day1 where
|
2024-06-04 14:31:23 +00:00
|
|
|
|
|
2024-06-11 02:56:36 +00:00
|
|
|
|
import Relation.Binary.PropositionalEquality as Eq
|
|
|
|
|
open import Agda.Builtin.Sigma
|
2024-06-07 14:22:38 +00:00
|
|
|
|
open import Data.Bool
|
2024-06-11 02:56:36 +00:00
|
|
|
|
open import Data.Empty
|
2024-06-07 14:22:38 +00:00
|
|
|
|
open import Data.Fin
|
2024-06-11 02:56:36 +00:00
|
|
|
|
open import Data.Maybe
|
|
|
|
|
open import Data.Nat
|
2024-06-07 14:22:38 +00:00
|
|
|
|
open import Data.Product
|
2024-06-11 02:56:36 +00:00
|
|
|
|
open import Data.Sum
|
|
|
|
|
open import Data.Unit
|
|
|
|
|
open import Relation.Nullary
|
|
|
|
|
|
2024-06-07 14:22:38 +00:00
|
|
|
|
open Eq using (_≡_; refl; trans; sym; cong; cong-app)
|
2024-06-11 02:56:36 +00:00
|
|
|
|
|
|
|
|
|
id : {A : Set} → A → A
|
|
|
|
|
id {A} x = x
|
2024-06-04 14:31:23 +00:00
|
|
|
|
|
2024-06-06 14:21:04 +00:00
|
|
|
|
data type : Set where
|
|
|
|
|
bool : type
|
2024-06-07 14:22:38 +00:00
|
|
|
|
_-→_ : type → type → type
|
2024-06-04 14:31:23 +00:00
|
|
|
|
|
2024-06-11 02:56:36 +00:00
|
|
|
|
data term : Set where
|
|
|
|
|
`_ : ℕ → term
|
|
|
|
|
`true : term
|
|
|
|
|
`false : term
|
|
|
|
|
`if_then_else_ : term → term → term → term
|
|
|
|
|
`λ[_]_ : (τ : type) → (e : term) → term
|
|
|
|
|
_∙_ : term → term → term
|
|
|
|
|
|
|
|
|
|
data ctx : Set where
|
|
|
|
|
nil : ctx
|
|
|
|
|
cons : ctx → type → ctx
|
|
|
|
|
|
|
|
|
|
lookup : ctx → ℕ → Maybe type
|
|
|
|
|
lookup nil _ = nothing
|
|
|
|
|
lookup (cons ctx₁ x) zero = just x
|
|
|
|
|
lookup (cons ctx₁ x) (suc n) = lookup ctx₁ n
|
|
|
|
|
|
|
|
|
|
data sub : Set where
|
|
|
|
|
nil : sub
|
|
|
|
|
cons : sub → term → sub
|
|
|
|
|
|
|
|
|
|
subst : term → term → term
|
|
|
|
|
subst (` zero) v = v
|
|
|
|
|
subst (` suc x) v = ` x
|
|
|
|
|
subst `true v = `true
|
|
|
|
|
subst `false v = `false
|
|
|
|
|
subst (`if x then x₁ else x₂) v = `if (subst x v) then (subst x₁ v) else (subst x₂ v)
|
|
|
|
|
subst (`λ[ τ ] x) v = `λ[ τ ] subst x v
|
|
|
|
|
subst (x ∙ x₁) v = (subst x v) ∙ (subst x₁ v)
|
|
|
|
|
|
|
|
|
|
data value-rel : type → term → Set where
|
|
|
|
|
v-`true : value-rel bool `true
|
|
|
|
|
v-`false : value-rel bool `false
|
|
|
|
|
v-`λ[_]_ : ∀ {τ e} → value-rel τ (`λ[ τ ] e)
|
|
|
|
|
|
|
|
|
|
data good-subst : ctx → sub → Set where
|
|
|
|
|
nil : good-subst nil nil
|
|
|
|
|
cons : ∀ {ctx τ γ v}
|
|
|
|
|
→ good-subst ctx γ
|
|
|
|
|
→ value-rel τ v
|
|
|
|
|
→ good-subst (cons ctx τ) γ
|
|
|
|
|
|
|
|
|
|
data step : term → term → Set where
|
|
|
|
|
step-if-1 : ∀ {e₁ e₂} → step (`if `true then e₁ else e₂) e₁
|
|
|
|
|
step-if-2 : ∀ {e₁ e₂} → step (`if `false then e₁ else e₂) e₁
|
|
|
|
|
step-`λ : ∀ {τ e v} → step ((`λ[ τ ] e) ∙ v) (subst e v)
|
|
|
|
|
|
|
|
|
|
data steps : ℕ → term → term → Set where
|
|
|
|
|
zero : ∀ {e} → steps zero e e
|
|
|
|
|
suc : ∀ {e e' e''} → (n : ℕ) → step e e' → steps n e' e'' → steps (suc n) e e''
|
|
|
|
|
|
|
|
|
|
data well-typed : ctx → term → type → Set where
|
|
|
|
|
type-`true : ∀ {ctx} → well-typed ctx `true bool
|
|
|
|
|
type-`false : ∀ {ctx} → well-typed ctx `false bool
|
|
|
|
|
type-`ifthenelse : ∀ {ctx e e₁ e₂ τ}
|
|
|
|
|
→ well-typed ctx e bool
|
|
|
|
|
→ well-typed ctx e₁ τ
|
|
|
|
|
→ well-typed ctx e₂ τ
|
|
|
|
|
→ well-typed ctx (`if e then e₁ else e₂) τ
|
|
|
|
|
type-`x : ∀ {ctx x}
|
|
|
|
|
→ (p : Is-just (lookup ctx x))
|
|
|
|
|
→ well-typed ctx (` x) (to-witness p)
|
|
|
|
|
type-`λ : ∀ {ctx τ τ₂ e}
|
|
|
|
|
→ well-typed (cons ctx τ) e τ₂
|
|
|
|
|
→ well-typed ctx (`λ[ τ ] e) (τ -→ τ₂)
|
|
|
|
|
type-∙ : ∀ {ctx τ₁ τ₂ e₁ e₂}
|
|
|
|
|
→ well-typed ctx e₁ (τ₁ -→ τ₂)
|
|
|
|
|
→ well-typed ctx e₂ τ₂
|
|
|
|
|
→ well-typed ctx (e₁ ∙ e₂) τ₂
|
|
|
|
|
|
|
|
|
|
_⊢_∶_ : (Γ : ctx) → (e : term) → (τ : type) → Set
|
|
|
|
|
Γ ⊢ e ∶ τ = (well-typed Γ e τ) × (∀ {n} → (e' : term) → steps n e e' → value-rel τ e' ⊎ ∃ λ e'' → step e' e'')
|
|
|
|
|
|
|
|
|
|
irreducible : term → Set
|
|
|
|
|
irreducible e = ¬ (∃ λ e' → step e e')
|
|
|
|
|
|
|
|
|
|
data term-relation : type → term → Set where
|
|
|
|
|
e-term : ∀ {τ e}
|
|
|
|
|
→ (∀ {n} → (e' : term) → steps n e e' → irreducible e' → value-rel τ e')
|
|
|
|
|
→ term-relation τ e
|
|
|
|
|
|
|
|
|
|
_⊨_∶_ : (Γ : ctx) → (e : term) → (τ : type) → Set
|
|
|
|
|
_⊨_∶_ Γ e τ = (γ : sub) → (good-subst Γ γ) → term-relation τ e
|
|
|
|
|
|
|
|
|
|
fundamental : ∀ {Γ e τ} → Γ ⊢ e ∶ τ → Γ ⊨ e ∶ τ
|
|
|
|
|
fundamental {Γ} {`true} {bool} type-sound γ good-sub = e-term λ e' steps irred →
|
|
|
|
|
[ id , (λ exists → ⊥-elim (irred exists)) ] (snd type-sound e' steps)
|
|
|
|
|
fundamental {Γ} {`false} {bool} type-sound γ good-sub = e-term λ e' steps irred →
|
|
|
|
|
[ id , (λ exists → ⊥-elim (irred exists)) ] (snd type-sound e' steps)
|
|
|
|
|
fundamental {Γ} {`λ[ τ ] e} {τ₁ -→ τ₂} type-sound γ good-sub = e-term f
|
|
|
|
|
where
|
|
|
|
|
f : {n : ℕ} (e' : term) → steps n (`λ[ τ ] e) e' → irreducible e' → value-rel (τ₁ -→ τ₂) e'
|
|
|
|
|
f e' steps irred = [ id , (λ exists → ⊥-elim (irred exists)) ] (snd type-sound e' steps)
|
|
|
|
|
|
|
|
|
|
fundamental {Γ} {`if e then e₁ else e₂} {τ} type-sound γ good-sub = e-term f
|
|
|
|
|
where
|
|
|
|
|
f : {n : ℕ} (e' : term) → steps n (`if e then e₁ else e₂) e' → irreducible e' → value-rel τ e'
|
|
|
|
|
f .(`if e then e₁ else e₂) zero irred =
|
|
|
|
|
let
|
|
|
|
|
ts : well-typed Γ (`if e then e₁ else e₂) τ
|
|
|
|
|
ts = fst type-sound
|
|
|
|
|
ts2 = snd type-sound
|
|
|
|
|
in ⊥-elim (irred {! !})
|
|
|
|
|
f e' (suc n x steps₁) irred = {! !}
|
|
|
|
|
fundamental {Γ} {` x} {τ} type-sound γ good-sub = {! !}
|
|
|
|
|
fundamental {Γ} {e ∙ e₁} {τ} type-sound γ good-sub = {! !}
|
|
|
|
|
|
|
|
|
|
-- fundamental {Γ} {`true} {τ -→ τ₁} type-sound γ good-sub = e-term f
|
|
|
|
|
-- where
|
|
|
|
|
-- el : ∀ {A} → well-typed Γ `true (τ -→ τ₁) → A
|
|
|
|
|
-- el ()
|
|
|
|
|
-- f : {n : ℕ} (e' : term) → steps n `true e' → irreducible e' → value-rel (τ -→ τ₁) e'
|
|
|
|
|
-- f e' steps irred = el (fst type-sound)
|
|
|
|
|
-- fundamental {Γ} {`false} {τ -→ τ₁} type-sound γ good-sub = e-term f
|
|
|
|
|
-- where
|
|
|
|
|
-- el : ∀ {A} → well-typed Γ `false (τ -→ τ₁) → A
|
|
|
|
|
-- el ()
|
|
|
|
|
-- f : {n : ℕ} (e' : term) → steps n `false e' → irreducible e' → value-rel (τ -→ τ₁) e'
|
|
|
|
|
-- f e' steps irred = el (fst type-sound)
|