diff --git a/out/Stlc.md b/out/Stlc.md index 2603053a..249d5263 100644 --- a/out/Stlc.md +++ b/out/Stlc.md @@ -4100,36 +4100,35 @@ Example terms. Construction of a type derivation is best done interactively. We start with the declaration: - `typing₁ : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹` - `typing₁ = ?` + typing₁ : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹 + typing₁ = ? -Typing control-L causes Agda to create a hole and tell us its expected type. +Typing C-L (control L) causes Agda to create a hole and tell us its +expected type. - `typing₁ = { }0` - `?0 : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹` + typing₁ = { }0 + ?0 : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹 -Now we fill in the hole, observing that the outermost term in `not` in a `λ`, -which is typed using `⇒-I`. The `⇒-I` rule in turn takes one argument, which -we again specify with a hole. +Now we fill in the hole by typing C-R (control R). Agda observes that +the outermost term in `not` in a `λ`, which is typed using `⇒-I`. The +`⇒-I` rule in turn takes one argument, which Agda leaves as a hole. - `typing₁ = ⇒-I { }0` - `?0 : ∅ , x ∶ 𝔹 ⊢ if var x then false else true ∶ 𝔹` + typing₁ = ⇒-I { }0 + ?0 : ∅ , x ∶ 𝔹 ⊢ if var x then false else true ∶ 𝔹 -Again we fill in the hole, observing that the outermost term is now -`if_then_else_`, which is typed using `𝔹-E`. The `𝔹-E` rule in turn takes -three arguments, which we again specify with holes. +Again we fill in the hole by typing C-R. Agda observes that the +outermost term is now `if_then_else_`, which is typed using `𝔹-E`. The +`𝔹-E` rule in turn takes three arguments, which Agda leaves as holes. - `typing₁ = ⇒-I (𝔹-E { }0 { }1 { }2)` - `?0 : ∅ , x ∶ 𝔹 ⊢ var x ∶ 𝔹` - `?1 : ∅ , x ∶ 𝔹 ⊢ false ∶ 𝔹` - `?2 : ∅ , x ∶ 𝔹 ⊢ true ∶ 𝔹` - -Again we fill in the three holes, observing that `var x`, `false`, and `true` -are typed using `Ax`, `𝔹-I₂`, and `𝔹-I₁` respectively. The `Ax` rule in turn -takes an argument, to show that `(∅ , x ∶ 𝔹) x = just 𝔹`, which can in turn -be computed with a hole. - -Filling in the three holes gives the derivation above. + typing₁ = ⇒-I (𝔹-E { }0 { }1 { }2) + ?0 : ∅ , x ∶ 𝔹 ⊢ var x ∶ + ?1 : ∅ , x ∶ 𝔹 ⊢ false ∶ 𝔹 + ?2 : ∅ , x ∶ 𝔹 ⊢ true ∶ 𝔹 +Again we fill in the three holes by typing C-R in each. Agda observes +that `var x`, `false`, and `true` are typed using `Ax`, `𝔹-I₂`, and +`𝔹-I₁` respectively. The `Ax` rule in turn takes an argument, to show +that `(∅ , x ∶ 𝔹) x = just 𝔹`, which can in turn be specified with a +hole. After filling in all holes, the term is as above. diff --git a/out/StlcProp.md b/out/StlcProp.md index 5dce75b7..cf7cf557 100644 --- a/out/StlcProp.md +++ b/out/StlcProp.md @@ -635,10 +635,8 @@ while for boolean types they are values `true` and `false`. >Ax ⊢x)()) {% raw %} -progressdata :Progress ∀: Term → Set where + steps : ∀ {M A} → ∅ ⊢ M ∶ A → Value M ⊎ ∃ λ N} → M ⟹ N → Progress M + done : ∀ {M} → Value M → Progress M + +progress : ∀ {M A} → ∅ ⊢ M ∶ A → Progress M {% endraw %} @@ -888,751 +982,485 @@ _Proof_: By induction on the derivation of `∅ ⊢ M ∶ A`. - If the last rule of the derivation is `⇒-E`, then the term has the form `L · M` for some `L` and `M`, where we know that `L` and `M` are also well typed in the empty context, giving us two induction - hypotheses. By the first induction hypothesis, either `L` is a - value or it can take a step. - - - If `L` is a value then consider `M`. By the second induction - hypothesis, either `M` is a value or it can take a step. - - - If `M` is a value, then since `L` is a value with a - function type we know it must be a lambda abstraction, - and hence `L · M` can take a step by `β⇒`. - - - If `M` can take a step, then so can `L · M` by `γ⇒₂`. + hypotheses. By the first induction hypothesis, either `L` + can take a step or is a value. - If `L` can take a step, then so can `L · M` by `γ⇒₁`. + - If `L` is a value then consider `M`. By the second induction + hypothesis, either `M` can take a step or is a value. + + - If `M` can take a step, then so can `L · M` by `γ⇒₂`. + + - If `M` is a value, then since `L` is a value with a + function type we know from the canonical forms lemma + that it must be a lambda abstraction, + and hence `L · M` can take a step by `β⇒`. + - If the last rule of the derivation is `𝔹-E`, then the term has the form `if L then M else N` where `L` has type `𝔹`. - By the induction hypothesis, either `L` is a value or it can - take a step. - - - If `L` is a value, then since it has type boolean it must be - either `true` or `false`. - - - If `L` is `true`, then then term steps by `β𝔹₁` - - - If `L` is `false`, then then term steps by `β𝔹₂` + By the induction hypothesis, either `L` can take a step or is + a value. - If `L` can take a step, then so can `if L then M else N` by `γ𝔹`. + - If `L` is a value, then since it has type boolean we know from + the canonical forms lemma that it must be either `true` or + `false`. + + - If `L` is `true`, then `if L then M else N` steps by `β𝔹₁` + + - If `L` is `false`, then `if L then M else N` steps by `β𝔹₂` + This completes the proof.
{% raw %} -progress (Ax ()) -progress (⇒-I ⊢N) = inj₁ value-λ -progress (⇒-E {Γ} {L} {M} {A} {B} ⊢L ⊢M) with progress ⊢L -... | inj₂ (L′ , L⟹L′) = inj₂ (L′ · M , γ⇒₁ L⟹L′) -... | inj₁ valueL with progress ⊢M -... | inj₂ (M′ , M⟹M′) =()) +progress inj₂ (L⇒-I · M′ , γ⇒₂ valueL M⟹M′⊢N) = done value-λ ...progress |(⇒-E inj₁ valueM with canonicalFormsLemma ⊢L valueL⊢M) with progress ⊢L ... | canonical-λsteps L⟹L′ = steps (γ⇒₁ {x} L⟹L′{.A}) +... {N}| {.B}done =valueL inj₂ ((Nwith [ x ∶= Mprogress ])⊢M , β⇒ valueM) progress... 𝔹-I₁| steps M⟹M′ = inj₁steps value-true(γ⇒₂ valueL M⟹M′) progress... | done 𝔹-I₂ =valueM inj₁with value-falsecanonicalFormsLemma ⊢L valueL progress (𝔹-E {Γ} {L}... {M}| {N}canonical-λ {A}= ⊢L ⊢M ⊢N)steps with(β⇒ progress ⊢LvalueM) ...progress |𝔹-I₁ inj₂= (L′done ,value-true +progress L⟹L′)𝔹-I₂ = inj₂ ((if L′ then Mdone else N) , γ𝔹 L⟹L′)value-false ...progress |(𝔹-E inj₁⊢L valueL⊢M ⊢N) with canonicalFormsLemmaprogress ⊢L +... | ⊢Lsteps valueLL⟹L′ = steps (γ𝔹 L⟹L′) ... | canonical-true = inj₂| (Mdone ,valueL β𝔹₁) -...with | canonical-false =canonicalFormsLemma inj₂⊢L (N ,valueL +... | canonical-true = steps β𝔹₁ +... | canonical-false = steps β𝔹₂) {% endraw %}@@ -1641,101 +1469,69 @@ Show that progress can also be proved by induction on terms instead of induction on typing derivations.
{% raw %} -postulate progress′ : ∀ M {M A} → ∅ ⊢ M ∶ A → ValueProgress M ⊎ ∃ λ N → M ⟹ N {% endraw %}@@ -1794,602 +1590,602 @@ For example: Formally:
{% raw %} -data _∈_ : Id → Term → Set where - free-var : ∀ {x} → x ∈ (var x) - free-λ : ∀ {x y A N} →where + free-var : y∀ ≢ {x} → x ∈ N → x ∈(var (λ[ y ∶ A ] Nx) free-·₁ free-λ : ∀ {x y A N} ∀→ {y ≢ x L M}→ →x x∈ ∈N L→ →x x∈ ∈ (Lλ[ ·y M)∶ A ] N) free-·₂free-·₁ : ∀ {x L∀ M}{x →L x ∈ M} → x ∈ (L → x ·∈ M)(L · M) free-if₁ : ∀free-·₂ : ∀ {x L Mx N}L M→} x→ ∈x L∈ →M x→ ∈x ∈ (if L then· M else N) free-if₂free-if₁ : ∀ {x L M :N} ∀→ {x L∈ ML N} → x ∈ M(if →L x ∈then (if L then M else N) free-if₃free-if₂ : ∀ {x L M :N} ∀→ {x L∈ M N} → x ∈ N(if →L x ∈then (if L then M else N) + free-if₃ : ∀ {x L M N} → x ∈ N → x ∈ (if L then M else N) {% endraw %}@@ -2397,72 +2193,72 @@ Formally: A term in which no variables appear free is said to be _closed_.
{% raw %} -closed : Term → Set closed M = ∀ {x} → ¬ (x ∈ M) {% endraw %}@@ -2483,115 +2279,115 @@ well typed in context `Γ`, then it must be the case that `Γ` assigns a type to `x`.
{% raw %} -freeLemma : ∀ {x M A Γ} → x ∈ M → Γ ⊢ M ∶ A → ∃ λ B → Γ x ≡ just B {% endraw %}@@ -2626,454 +2422,454 @@ _Proof_: We show, by induction on the proof that `x` appears `_≟_`, noting that `x` and `y` are different variables.
{% raw %} -freeLemma free-var (Ax Γx≡justA) = (_ , Γx≡justA) -freeLemma (free-·₁ x∈L) (⇒-E ⊢L ⊢M) = freeLemma(_ x∈L, ⊢LΓx≡justA) freeLemma (free-·₂free-·₁ x∈Mx∈L) (⇒-E ⊢L ⊢M) = ⊢M) = freeLemma x∈Mx∈L ⊢M⊢L freeLemma (free-if₁free-·₂ x∈Lx∈M) (𝔹-E⇒-E ⊢L ⊢M) ⊢L ⊢M ⊢N) = freeLemma freeLemmax∈M x∈L ⊢L⊢M freeLemma (free-if₁ x∈L) (free-if₂𝔹-E x∈M)⊢L (𝔹-E⊢M ⊢N) ⊢L ⊢M ⊢N) = freeLemma freeLemmax∈L x∈M ⊢M⊢L freeLemma (free-if₂ x∈M) (free-if₃𝔹-E x∈N)⊢L (𝔹-E⊢M ⊢N) ⊢L ⊢M ⊢N) = freeLemma freeLemmax∈M x∈N ⊢N⊢M freeLemma (free-if₃ x∈N) (free-λ𝔹-E {x}⊢L {y}⊢M y≢x⊢N) = freeLemma x∈N) (⇒-I ⊢N) with +freeLemma x∈N(free-λ ⊢N -... | Γx=justC with y ≟ {x -...} |{y} y≢x x∈N) (⇒-I ⊢N) with yes y≡x =freeLemma ⊥-elimx∈N ⊢N (y≢x y≡x) ... | no _Γx=justC with y = ≟ x +... | yes y≡x = ⊥-elim (y≢x y≡x) +... | no _ = Γx=justC {% endraw %}@@ -3088,362 +2884,362 @@ the empty context is closed (it has no free variables). #### Exercise: 2 stars, optional (∅⊢-closed)
{% raw %} -postulate ∅⊢-closed : ∀ {M A} → ∅ ⊢ M ∶ A → closed M {% endraw %}
{% raw %} -contradiction : ∀ {X : Set} → ∀ {x : X} → ¬ (_≡_ {A = Maybe X} (just: ∀ {X x): nothingSet} → )∀ -contradiction () - -∅⊢-closed′{x : X} → ¬ (_≡_ {A = Maybe X} ∀(just {Mx) A} → ∅ ⊢ M ∶nothing) A → closed M contradiction () + +∅⊢-closed′ : ∀ {M} {A} ⊢M {xA} x∈M→ ∅ with⊢ freeLemmaM x∈M∶ A → closed M ⊢M ...∅⊢-closed′ |{M} (B{A} ,⊢M ∅x≡justB) {x=} contradictionx∈M (transwith (symfreeLemma x∈M ⊢M +... ∅x≡justB)| (apply-∅B , ∅x≡justB) = contradiction (trans (sym ∅x≡justB) (apply-∅ x)) {% endraw %}@@ -3457,144 +3253,144 @@ that appear free in `M`. In fact, this is the only condition that is needed.
{% raw %} -contextLemma : ∀ {Γ Γ′ M A} - → (∀ {x} → x ∈ M → Γ x ≡ Γ′ x) - →∀ {Γ ⊢ Γ′ M ∶ A} → (∀ {x} → x Γ′∈ ⊢ M → Γ x ≡ Γ′ x) + → Γ ⊢ M ∶ A + → Γ′ ⊢ M ∶ A {% endraw %}@@ -3642,604 +3438,604 @@ _Proof_: By induction on the derivation of hence the desired result follows from the induction hypotheses.
{% raw %} -contextLemma Γ~Γ′ (Ax Γx≡justA) rewrite (Γ~Γ′ free-var) = Ax Γx≡justA -contextLemma {Γ} {Γ′} {λ[ x ∶ A ] N}= Γ~Γ′Ax (⇒-I ⊢N) = ⇒-I (Γx≡justA +contextLemma {Γ} {Γ′} Γx~Γ′x ⊢N) - where - Γx~Γ′x : ∀ {y}λ[ →x y∶ ∈A ] N} →Γ~Γ′ (Γ⇒-I ⊢N) = ⇒-I (contextLemma , x ∶ A) y ≡ (Γ′ , x ∶ A) y - Γx~Γ′x {y} y∈N with x ≟ y⊢N) ...where + Γx~Γ′x : ∀ {y} → y ∈ N → (Γ , x ∶ A) y ≡ |(Γ′ yes, reflx ∶ A=) refly ...Γx~Γ′x |{y} no x≢y =y∈N with Γ~Γ′x (free-λ≟ x≢y y∈N) -contextLemma Γ~Γ′ (⇒-E ⊢L ⊢M) = ⇒-E (contextLemma (Γ~Γ′ ∘ free-·₁) ⊢L) (contextLemma (Γ~Γ′ ∘ free-·₂) ⊢M) -contextLemma Γ~Γ′ 𝔹-I₁ = 𝔹-I₁ -contextLemma Γ~Γ′ 𝔹-I₂ = 𝔹-I₂ -contextLemma Γ~Γ′ (𝔹-E ⊢L ⊢M ⊢N)y ... | yes refl = refl + ... | no x≢y = Γ~Γ′ (free-λ x≢y y∈N) +contextLemma Γ~Γ′ (⇒-E ⊢L ⊢M) = ⇒-E (contextLemma (Γ~Γ′ ∘ free-·₁) ⊢L) (contextLemma (Γ~Γ′ ∘ free-·₂) ⊢M) +contextLemma Γ~Γ′ 𝔹-I₁ = 𝔹-I₁ +contextLemma Γ~Γ′ 𝔹-I₂ = 𝔹-I₂ +contextLemma Γ~Γ′ (𝔹-E (contextLemma (Γ~Γ′ ∘ free-if₁) ⊢L) (contextLemma (Γ~Γ′ ∘ free-if₂) ⊢M) (contextLemma (Γ~Γ′ ∘ free-if₃) ⊢N) + = 𝔹-E (contextLemma (Γ~Γ′ ∘ free-if₁) ⊢L) (contextLemma (Γ~Γ′ ∘ free-if₂) ⊢M) (contextLemma (Γ~Γ′ ∘ free-if₃) ⊢N) {- replaceCtxt f (var x x∶A ) rewrite f var = var x x∶A @@ -4280,162 +4076,162 @@ _Lemma_: If `Γ , x ∶ A ⊢ N ∶ B` and `∅ ⊢ V ∶ A`, then `Γ ⊢ (N [ x ∶= V ]) ∶ B`.{% raw %} -preservation-[∶=] : ∀ {Γ x A N B V} - → (Γ , x ∶ A) ⊢: N∀ ∶ B - → ∅ ⊢ V ∶ A - → {Γ ⊢ (N [ x ∶=A N B V} + → ])(Γ , x ∶ A) ⊢ N ∶ B + → ∅ ⊢ V ∶ A + → Γ ⊢ (N [ x ∶= V ]) ∶ B {% endraw %}@@ -4505,1337 +4301,1337 @@ _Proof_: We show, by induction on `N`, that for all `A` and We need a couple of lemmas. A closed term can be weakened to any context, and just is injective.{% raw %} -weaken-closed : ∀ {V A Γ} → ∅ ⊢ V ∶ A → Γ ⊢ V ∶ A -weaken-closed {V} {A} {Γ} ⊢V→ ∅ =⊢ contextLemmaV Γ~Γ′∶ A → Γ ⊢ V ∶ A +weaken-closed ⊢V - where - Γ~Γ′{V} : {A∀} {xΓ} →⊢V x= ∈ V → ∅ x ≡ ΓcontextLemma xΓ~Γ′ ⊢V where + Γ~Γ′ : ∀ {x} x∈V = ⊥-elim (x∉V x∈V) - where - x∉V→ :x ¬∈ V → ∅ x ≡ Γ x + Γ~Γ′ ({x} ∈ V)x∈V = ⊥-elim (x∉V x∈V) x∉V = ∅⊢-closed ⊢V {x}where - -just-injectivex∉V : ∀¬ (x ∈ V{X : Set}) + x∉V {x= y : X} →∅⊢-closed _≡_⊢V {A = Maybe X} (just x} + +just-injective : )∀ (just{X y): →Set} {x ≡y y: -just-injective reflX} → _≡_ {A = Maybe =X} (just x) (just y) → x ≡ y +just-injective refl = refl {% endraw %}
{% raw %} -preservation-[∶=] {_} {x} (Ax {_} {x′} [Γ,x∶A]x′≡B) ⊢V with x ≟ x′ - ...|{_} yes{x} x≡x′(Ax rewrite{_} just-injective{x′} [Γ,x∶A]x′≡B) ⊢V [Γ,x∶A]x′≡B = weaken-closedwith ⊢V -...| no x≢x′ = Ax [Γ,x∶A]x′≡B -preservation-[∶=] {Γ} {x} {A}≟ {λ[ x′ +...| yes x≡x′ rewrite just-injective [Γ,x∶A]x′≡B = weaken-closed ⊢V +...| no x≢x′ = Ax ∶ A′ ] N′} {.A′ ⇒ B′} {V} (⇒-I ⊢N′) ⊢V with x ≟ x′[Γ,x∶A]x′≡B ...|preservation-[∶=] {Γ} {x} {A} {λ[ x′ ∶ A′ ] yesN′} x≡x′{.A′ rewrite⇒ x≡x′B′} ={V} (⇒-I contextLemma⊢N′) ⊢V with Γ′~Γx (⇒-I≟ ⊢N′)x′ - where - Γ′~Γ...| :yes ∀x≡x′ rewrite x≡x′ {y}= →contextLemma yΓ′~Γ ∈ (λ[ x′ ∶ A′ ] N′)⇒-I ⊢N′→ (Γ , x′ ∶ A) y + where + Γ′~Γ ≡: Γ∀ y - Γ′~Γ {y} → y ∈ (free-λλ[ x′≢y y∈N′) with x′ ∶ A′ ] N′) → (Γ , x′ ∶ ≟A) y ≡ Γ y ...|Γ′~Γ yes{y} x′≡y = ⊥-elim(free-λ (x′≢y x′≡yy∈N′) - ...| with x′ no _ =≟ refly -...| noyes x′≡y x≢x′ = ⇒-I⊥-elim ⊢N′V(x′≢y x′≡y) where - x′x⊢N′...| no _ := Γrefl +...| ,no x≢x′ x′= ∶ A′ , x ∶ A ⊢ N′⇒-I ∶ B′⊢N′V where + x′x⊢N′ rewrite update-permute: Γ x, A x′ ∶ A′ , x ∶ A ⊢ A′N′ x≢x′∶ B′ = ⊢N′ ⊢N′Vx′x⊢N′ :rewrite (update-permute Γ , x′ ∶ A′) ⊢ N′ [x xA ∶=x′ VA′ ]x≢x′ ∶= B′⊢N′ ⊢N′V =: preservation-[∶=] x′x⊢N′ ⊢V -preservation-[∶=] (⇒-EΓ ⊢L, ⊢Mx′ ∶ A′) ⊢V⊢ N′ [ x ∶= V ] ∶ B′ + ⊢N′V = ⇒-Epreservation-[∶=] (preservation-[∶=]x′x⊢N′ ⊢L ⊢V) ( +preservation-[∶=] (⇒-E ⊢L ⊢M) ⊢V = ⊢M⇒-E ⊢V) -(preservation-[∶=] 𝔹-I₁⊢L ⊢V) ⊢V = 𝔹-I₁ -preservation-[∶=] 𝔹-I₂ ⊢V = 𝔹-I₂ -preservation-[∶=] (𝔹-Epreservation-[∶=] ⊢L ⊢M ⊢N) ⊢V =) - 𝔹-Epreservation-[∶=] 𝔹-I₁ ⊢V = 𝔹-I₁ +preservation-[∶=] 𝔹-I₂ ⊢V = (𝔹-I₂ +preservation-[∶=] (𝔹-E ⊢L ⊢V) (preservation-[∶=] ⊢M ⊢V⊢N) ⊢V = + 𝔹-E (preservation-[∶=] ⊢L ⊢N ⊢V) (preservation-[∶=] ⊢M ⊢V) (preservation-[∶=] ⊢N ⊢V) {% endraw %}@@ -5849,95 +5645,95 @@ is also a closed term with type `A`. In other words, small-step reduction preserves types.
{% raw %} -preservation : ∀ {M N A} → ∅ ⊢ M ∶ A → M ⟹ N → ∅ ⊢ N ∶ A {% endraw %}@@ -5974,435 +5770,435 @@ _Proof_: By induction on the derivation of `\vdash t : T`. follows directly from the induction hypothesis.
{% raw %} -preservation (Ax x₁) () -preservation (⇒-I ⊢N) () -preservation (⇒-EAx (⇒-I ⊢Nx₁) ⊢V) (β⇒ valueV) = preservation-[∶=] ⊢N ⊢V() preservation (⇒-E⇒-I ⊢L ⊢M⊢N) () +preservation (⇒-E (⇒-I ⊢N) ⊢V) (β⇒ valueV) = (γ⇒₁preservation-[∶=] L⟹L′)⊢N with preservation ⊢L L⟹L′⊢V preservation ...|(⇒-E ⊢L′ = ⇒-E ⊢L′ ⊢M -preservation (⇒-E ⊢L ⊢M) (γ⇒₂γ⇒₁ valueL M⟹M′L⟹L′) with preservation ⊢M⊢L M⟹M′L⟹L′ ...| ⊢L′ = ⇒-E ⊢L′ ⊢M +preservation (⇒-E ⊢L ⊢M) ⊢M′ = ⇒-E ⊢L ⊢M′ -preservation 𝔹-I₁ () -preservation 𝔹-I₂ () -preservation (𝔹-Eγ⇒₂ 𝔹-I₁valueL ⊢M ⊢NM⟹M′) β𝔹₁with =preservation ⊢M M⟹M′ ...| ⊢M′ = ⇒-E ⊢L ⊢M′ +preservation (𝔹-E𝔹-I₁ () +preservation 𝔹-I₂ () +preservation (𝔹-E 𝔹-I₁ ⊢M ⊢N) β𝔹₂ = ⊢N) β𝔹₁ = ⊢M preservation (𝔹-E ⊢L ⊢M ⊢N) (γ𝔹𝔹-I₂ L⟹L′⊢M ⊢N) withβ𝔹₂ preservation= ⊢L L⟹L′⊢N ...|preservation ⊢L′(𝔹-E ⊢L ⊢M =⊢N) 𝔹-E(γ𝔹 L⟹L′) with preservation ⊢L L⟹L′ +...| ⊢L′ = 𝔹-E ⊢L′ ⊢M ⊢N {% endraw %}diff --git a/src/Stlc.lagda b/src/Stlc.lagda index 0f5e39aa..53ccdfeb 100644 --- a/src/Stlc.lagda +++ b/src/Stlc.lagda @@ -206,36 +206,35 @@ typing₂ = ⇒-I (⇒-I (⇒-E (Ax refl) (⇒-E (Ax refl) (Ax refl)))) Construction of a type derivation is best done interactively. We start with the declaration: - `typing₁ : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹` - `typing₁ = ?` + typing₁ : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹 + typing₁ = ? -Typing control-L causes Agda to create a hole and tell us its expected type. +Typing C-L (control L) causes Agda to create a hole and tell us its +expected type. - `typing₁ = { }0` - `?0 : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹` + typing₁ = { }0 + ?0 : ∅ ⊢ not ∶ 𝔹 ⇒ 𝔹 -Now we fill in the hole, observing that the outermost term in `not` in a `λ`, -which is typed using `⇒-I`. The `⇒-I` rule in turn takes one argument, which -we again specify with a hole. +Now we fill in the hole by typing C-R (control R). Agda observes that +the outermost term in `not` in a `λ`, which is typed using `⇒-I`. The +`⇒-I` rule in turn takes one argument, which Agda leaves as a hole. - `typing₁ = ⇒-I { }0` - `?0 : ∅ , x ∶ 𝔹 ⊢ if var x then false else true ∶ 𝔹` + typing₁ = ⇒-I { }0 + ?0 : ∅ , x ∶ 𝔹 ⊢ if var x then false else true ∶ 𝔹 -Again we fill in the hole, observing that the outermost term is now -`if_then_else_`, which is typed using `𝔹-E`. The `𝔹-E` rule in turn takes -three arguments, which we again specify with holes. +Again we fill in the hole by typing C-R. Agda observes that the +outermost term is now `if_then_else_`, which is typed using `𝔹-E`. The +`𝔹-E` rule in turn takes three arguments, which Agda leaves as holes. - `typing₁ = ⇒-I (𝔹-E { }0 { }1 { }2)` - `?0 : ∅ , x ∶ 𝔹 ⊢ var x ∶ 𝔹` - `?1 : ∅ , x ∶ 𝔹 ⊢ false ∶ 𝔹` - `?2 : ∅ , x ∶ 𝔹 ⊢ true ∶ 𝔹` - -Again we fill in the three holes, observing that `var x`, `false`, and `true` -are typed using `Ax`, `𝔹-I₂`, and `𝔹-I₁` respectively. The `Ax` rule in turn -takes an argument, to show that `(∅ , x ∶ 𝔹) x = just 𝔹`, which can in turn -be computed with a hole. - -Filling in the three holes gives the derivation above. + typing₁ = ⇒-I (𝔹-E { }0 { }1 { }2) + ?0 : ∅ , x ∶ 𝔹 ⊢ var x ∶ + ?1 : ∅ , x ∶ 𝔹 ⊢ false ∶ 𝔹 + ?2 : ∅ , x ∶ 𝔹 ⊢ true ∶ 𝔹 +Again we fill in the three holes by typing C-R in each. Agda observes +that `var x`, `false`, and `true` are typed using `Ax`, `𝔹-I₂`, and +`𝔹-I₁` respectively. The `Ax` rule in turn takes an argument, to show +that `(∅ , x ∶ 𝔹) x = just 𝔹`, which can in turn be specified with a +hole. After filling in all holes, the term is as above. diff --git a/src/StlcProp.lagda b/src/StlcProp.lagda index 05538372..17604c25 100644 --- a/src/StlcProp.lagda +++ b/src/StlcProp.lagda @@ -39,7 +39,7 @@ data canonical_for_ : Term → Type → Set where canonical-false : canonical false for 𝔹 canonicalFormsLemma : ∀ {L A} → ∅ ⊢ L ∶ A → Value L → canonical L for A -canonicalFormsLemma (Ax ⊢x) () +canonicalFormsLemma (Ax ()) () canonicalFormsLemma (⇒-I ⊢N) value-λ = canonical-λ canonicalFormsLemma (⇒-E ⊢L ⊢M) () canonicalFormsLemma 𝔹-I₁ value-true = canonical-true @@ -50,11 +50,15 @@ canonicalFormsLemma (𝔹-E ⊢L ⊢M ⊢N) () ## Progress As before, the _progress_ theorem tells us that closed, well-typed -terms are not stuck: either a well-typed term is a value, or it -can take a reduction step. +terms are not stuck: either a well-typed term can take a reduction +step or it is a value. \begin{code} -progress : ∀ {M A} → ∅ ⊢ M ∶ A → Value M ⊎ ∃ λ N → M ⟹ N +data Progress : Term → Set where + steps : ∀ {M N} → M ⟹ N → Progress M + done : ∀ {M} → Value M → Progress M + +progress : ∀ {M A} → ∅ ⊢ M ∶ A → Progress M \end{code} The proof is a relatively @@ -74,52 +78,54 @@ _Proof_: By induction on the derivation of `∅ ⊢ M ∶ A`. - If the last rule of the derivation is `⇒-E`, then the term has the form `L · M` for some `L` and `M`, where we know that `L` and `M` are also well typed in the empty context, giving us two induction - hypotheses. By the first induction hypothesis, either `L` is a - value or it can take a step. - - - If `L` is a value then consider `M`. By the second induction - hypothesis, either `M` is a value or it can take a step. - - - If `M` is a value, then since `L` is a value with a - function type we know it must be a lambda abstraction, - and hence `L · M` can take a step by `β⇒`. - - - If `M` can take a step, then so can `L · M` by `γ⇒₂`. + hypotheses. By the first induction hypothesis, either `L` + can take a step or is a value. - If `L` can take a step, then so can `L · M` by `γ⇒₁`. + - If `L` is a value then consider `M`. By the second induction + hypothesis, either `M` can take a step or is a value. + + - If `M` can take a step, then so can `L · M` by `γ⇒₂`. + + - If `M` is a value, then since `L` is a value with a + function type we know from the canonical forms lemma + that it must be a lambda abstraction, + and hence `L · M` can take a step by `β⇒`. + - If the last rule of the derivation is `𝔹-E`, then the term has the form `if L then M else N` where `L` has type `𝔹`. - By the induction hypothesis, either `L` is a value or it can - take a step. - - - If `L` is a value, then since it has type boolean it must be - either `true` or `false`. - - - If `L` is `true`, then then term steps by `β𝔹₁` - - - If `L` is `false`, then then term steps by `β𝔹₂` + By the induction hypothesis, either `L` can take a step or is + a value. - If `L` can take a step, then so can `if L then M else N` by `γ𝔹`. + - If `L` is a value, then since it has type boolean we know from + the canonical forms lemma that it must be either `true` or + `false`. + + - If `L` is `true`, then `if L then M else N` steps by `β𝔹₁` + + - If `L` is `false`, then `if L then M else N` steps by `β𝔹₂` + This completes the proof. \begin{code} progress (Ax ()) -progress (⇒-I ⊢N) = inj₁ value-λ -progress (⇒-E {Γ} {L} {M} {A} {B} ⊢L ⊢M) with progress ⊢L -... | inj₂ (L′ , L⟹L′) = inj₂ (L′ · M , γ⇒₁ L⟹L′) -... | inj₁ valueL with progress ⊢M -... | inj₂ (M′ , M⟹M′) = inj₂ (L · M′ , γ⇒₂ valueL M⟹M′) -... | inj₁ valueM with canonicalFormsLemma ⊢L valueL -... | canonical-λ {x} {.A} {N} {.B} = inj₂ ((N [ x ∶= M ]) , β⇒ valueM) -progress 𝔹-I₁ = inj₁ value-true -progress 𝔹-I₂ = inj₁ value-false -progress (𝔹-E {Γ} {L} {M} {N} {A} ⊢L ⊢M ⊢N) with progress ⊢L -... | inj₂ (L′ , L⟹L′) = inj₂ ((if L′ then M else N) , γ𝔹 L⟹L′) -... | inj₁ valueL with canonicalFormsLemma ⊢L valueL -... | canonical-true = inj₂ (M , β𝔹₁) -... | canonical-false = inj₂ (N , β𝔹₂) +progress (⇒-I ⊢N) = done value-λ +progress (⇒-E ⊢L ⊢M) with progress ⊢L +... | steps L⟹L′ = steps (γ⇒₁ L⟹L′) +... | done valueL with progress ⊢M +... | steps M⟹M′ = steps (γ⇒₂ valueL M⟹M′) +... | done valueM with canonicalFormsLemma ⊢L valueL +... | canonical-λ = steps (β⇒ valueM) +progress 𝔹-I₁ = done value-true +progress 𝔹-I₂ = done value-false +progress (𝔹-E ⊢L ⊢M ⊢N) with progress ⊢L +... | steps L⟹L′ = steps (γ𝔹 L⟹L′) +... | done valueL with canonicalFormsLemma ⊢L valueL +... | canonical-true = steps β𝔹₁ +... | canonical-false = steps β𝔹₂ \end{code} #### Exercise: 3 stars, optional (progress_from_term_ind) @@ -128,7 +134,7 @@ instead of induction on typing derivations. \begin{code} postulate - progress′ : ∀ {M A} → ∅ ⊢ M ∶ A → Value M ⊎ ∃ λ N → M ⟹ N + progress′ : ∀ M {A} → ∅ ⊢ M ∶ A → Progress M \end{code} ## Preservation