complete day 1

This commit is contained in:
Michael Zhang 2024-06-11 09:09:51 -04:00
parent be266049f0
commit a01be24558

View file

@ -1,7 +1,6 @@
{-# OPTIONS --prop #-}
module Ahmed.Day1 where
import Relation.Binary.PropositionalEquality as Eq
open import Agda.Builtin.Sigma
open import Data.Bool
open import Data.Empty
@ -10,11 +9,8 @@ open import Data.Maybe
open import Data.Nat
open import Data.Product
open import Data.Sum
open import Data.Unit
open import Relation.Nullary
open Eq using (_≡_; refl; trans; sym; cong; cong-app)
id : {A : Set} A A
id {A} x = x
@ -66,7 +62,7 @@ data good-subst : ctx → sub → Set where
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-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
@ -107,37 +103,7 @@ _⊨__ : (Γ : 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
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)
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)