diff --git a/out/Maps.md b/out/Maps.md index 11087b7e..db033d23 100644 --- a/out/Maps.md +++ b/out/Maps.md @@ -11,7 +11,7 @@ a nice case study using ideas we've seen in previous chapters, including building data structures out of higher-order functions (from [Basics]({{ "Basics" | relative_url }}) and [Poly]({{ "Poly" | relative_url }}) and the use of reflection to streamline proofs (from [IndProp]({{ "IndProp" | relative_url -}})). +}})). We'll define two flavors of maps: _total_ maps, which include a "default" element to be returned when a key being looked up @@ -27,216 +27,224 @@ Unlike the chapters we have seen so far, this one does not import the chapter before it (and, transitively, all the earlier chapters). Instead, in this chapter and from now, on we're going to import the definitions and theorems we need -directly from Agda's standard library stuff. You should not notice +directly from Agda's standard library. You should not notice much difference, though, because we've been careful to name our own definitions and theorems the same as their counterparts in the standard library, wherever they overlap. -
{% raw %} - + +open import import FunctionData.Nat using (_∘_ℕ) open import Data.Bool using (Bool; true; false) -open import Data.Empty using (⊥; ⊥-elim) open import Data.Maybe using (Maybe; just; nothing) +open import Data.String using (String) +open just; nothing) -open import Data.Nat using (ℕ) -open import Relation.Nullary using (¬_; Dec; yes; no) open import Relation.Binary.PropositionalEquality -{% endraw %}+ using (_≡_; refl; _≢_; trans; sym) + + Documentation for the standard library can be found at
{% raw %} - + +data Id : Set where id : ℕString → Id -{% endraw %}-
{% raw %} -_≟_ + +We recall a standard fact of logic. + ++ +Using the above, we can decide equality of two identifiers +by deciding equality on the underlying strings. + ++ +contrapositive : (x∀ y{ℓ₁ ℓ₂} {P : Id)Set ℓ₁} {Q : Set ℓ₂} → Dec (xP ≡→ yQ) -id x→ ≟(¬ idQ → ¬ y with x Data.Nat.≟ yP) idcontrapositive p→q ¬q xp≟= id¬q (p→q p) + +
+ +_≟_ : (x y |: yesId) x=y→ rewriteDec x=y(x =≡ yes refly) id x ≟ id y with x Data.String.≟ y +id x ≟ id y | noyes refl x≠y = yes refl +id x ≟ id y | no x≢y = no (x≠ycontrapositive ∘ id-inj x≢y) where id-inj : ∀ {x y} → id x ≡ id y → x ≡ y id-inj refl = refl -{% endraw %}+ + + +We define some identifiers for future use. + +
+ +x y z : Id +x = id "x" +y = id "y" +z = id "z" + ++ +## Extensionality ## Total Maps @@ -599,1038 +839,1669 @@ We build partial maps in two steps. First, we define a type of _total maps_ that return a default value when we look up a key that is not present in the map. -
{% raw %} - + +TotalMap : Set → Set TotalMap A = Id → A -{% endraw %}+ + Intuitively, a total map over anfi element type $$A$$ _is_ just a function that can be used to look up ids, yielding $$A$$s. -
{% raw %} - + +module TotalMap where -{% endraw %}-The function `empty` yields an empty total map, given a + + +The function `always` yields a total map given a default element; this map always returns the default element when applied to any id. -
{% raw %} - empty : ∀ {A} → A → TotalMap A - empty x = λ _ → x -{% endraw %}+
-More interesting is the `update` function, which (as before) takes -a map $$m$$, a key $$x$$, and a value $$v$$ and returns a new map that -takes $$x$$ to $$v$$ and takes every other key to whatever $$m$$ does. - -This definition is a nice example of higher-order programming. -The `update` function takes a _function_ $$m$$ and yields a new -function $$\lambda x'.\vdots$$ that behaves like the desired map. +The update function takes a _function_ $$ρ$$ and yields a new +function that behaves like the desired map. -For example, we can build a map taking ids to bools, where `id -3` is mapped to `true` and every other key is mapped to `false`, -like this: +We define handy abbreviations for updating a map two, three, or four times. -{% raw %} - update : ∀ {A} → TotalMap A -> Id -> A -> TotalMap A - update m x v yalways with: x∀ ≟{A} y - ...→ |A yes→ TotalMap x=y = vA ...always |v no x≠yx = mv + ++ +More interesting is the update function, which (as before) takes +a map $$ρ$$, a key $$x$$, and a value $$v$$ and returns a new map that +takes $$x$$ to $$v$$ and takes every other key to whatever $$ρ$$ does. + ++ + _,_↦_ : ∀ {A} → TotalMap A → Id → A → TotalMap A + (ρ , x ↦ v) y with x ≟ y + ... | yes x=y = v + ... | no x≠y = ρ y -{% endraw %}+ +
{% raw %} - examplemap +Wen: you don't actually need to define these, you can simply declare `_,_↦_` to +be a left-associative infix operator with an `infixl` statement, and then you'll +be able to just evaluate `M , x ↦ y , z ↦ w` as `(M , x ↦ y) , z ↦ w`. + + ++ +For example, we can build a map taking ids to naturals, where $$x$$ +maps to 42, $$y$$ maps to 69, and every other key maps to 0, as follows: + ++ + _,_↦_,_↦_ : ∀ {A} → TotalMap BoolA → Id → A → Id → A → TotalMap A examplemapρ , x₁ ↦ v₁ , x₂ ↦ v₂ = update (updateρ (empty, falsex₁ ↦ v₁), (idx₂ 1)↦ false) (id 3) truev₂ -{% endraw %}+ + _,_↦_,_↦_,_↦_ : ∀ {A} → TotalMap A → Id → A → Id → A → Id → A → TotalMap A + ρ , x₁ ↦ v₁ , x₂ ↦ v₂ , x₃ ↦ v₃ = ((ρ , x₁ ↦ v₁), x₂ ↦ v₂), x₃ ↦ v₃ + + _,_↦_,_↦_,_↦_,_↦_ : ∀ {A} → TotalMap A → Id → A → Id → A → Id → A → Id → A → TotalMap A + ρ , x₁ ↦ v₁ , x₂ ↦ v₂ , x₃ ↦ v₃ , x₄ ↦ v₄ = (((ρ , x₁ ↦ v₁), x₂ ↦ v₂), x₃ ↦ v₃), x₄ ↦ v₄ + +
+ + ρ₀ : TotalMap ℕ + ρ₀ = always 0 , x ↦ 42 , y ↦ 69 + +This completes the definition of total maps. Note that we don't need to define a `find` operation because it is just function application! -
{% raw %} - test_examplemap0 + + test₁ : examplemapρ₀ (idx ≡ 42 + test₁ = refl + + test₂ : ρ₀ y ≡ 69 + test₂ = refl + + test₃ : ρ₀ z ≡ 0) ≡ false test_examplemap0test₃ = refl - test_examplemap1 : examplemap (id 1) ≡ false - test_examplemap1 = refl - - test_examplemap2 : examplemap (id 2) ≡ false - test_examplemap2 = refl - - test_examplemap3 : examplemap (id 3) ≡ true - test_examplemap3 = refl -{% endraw %}+ To use maps in later chapters, we'll need several fundamental facts about how they behave. Even if you don't work the following -exercises, make sure you thoroughly understand the statements of -the lemmas! (Some of the proofs require the functional -extensionality axiom, which is discussed in the [Logic] -chapter and included in the Agda standard library.) +exercises, make sure you understand the statements of +the lemmas! -#### Exercise: 1 star, optional (apply-empty) -First, the empty map returns its default element for all keys: +#### Exercise: 1 star, optional (apply-always) +The `always` map returns its default element for all keys: -
{% raw %} - + + postulate apply-emptyapply-always : ∀ {A x v} (v : A) (x : Id) → emptyalways {A} v x ≡ v -{% endraw %}+ +
{% raw %} - apply-empty′ + + apply-always′ : ∀ {A x v} (v : A) (x : Id) → emptyalways {A} v x ≡ v apply-empty′apply-always′ v x = refl -{% endraw %}+ +
{% raw %} - + + postulate update-eq : ∀ {A v} (mρ : TotalMap A) (x : Id) (v : A) → (updateρ m, x ↦ v) x ≡ v -{% endraw %}+ +
{% raw %} - + + update-eq′ : ∀ {A v} (mρ : TotalMap A) (x : Id) (v : A) → (updateρ m, x ↦ v) x ≡ v update-eq′ mρ x v with x ≟ x ... | yes x=xx≡x = refl ... | no x≠xx≢x = ⊥-elim (x≠xx≢x refl) -{% endraw %}+ +
{% raw %} - + + update-neq : ∀ {A v} (mρ : TotalMap A) (x y : Id) - → x ≢ y → (update m x v) y ≡ m y - update-neq m x y x≠y with x ≟ y - ... | yes x=y = ⊥-elim (x≠y x=y) - ... | no _ = refl -{% endraw %}- -#### Exercise: 2 stars, optional (update-shadow) -If we update a map $$m$$ at a key $$x$$ with a value $$v_1$$ and then -update again with the same key $$x$$ and another value $$v_2$$, the -resulting map behaves the same (gives the same result when applied -to any key) as the simpler map obtained by performing just -the second `update` on $$m$$: - -
{% raw %} - postulate - update-shadow : ∀ {A v1 v2} (m : TotalMap A) (x y : Id) - → (update (update m x v1) x v2) y ≡ (update m x v2) y -{% endraw %}- -
{% raw %} - update-shadow′ : ∀ {A v1 v2} (m : TotalMap A) (x y : IdA) - → (x : Id) (v : A) (y (update: (update m x v1Id) + → x v2)≢ y → (ρ ≡, (update m x v2↦ v) y ≡ ρ y update-shadow′update-neq mρ x v y - x≢y with x ≟ y ... | yes x=yx≡y = ⊥-elim refl(x≢y x≡y) ... | no x≠y | no _ = update-neq m x y x≠yrefl -{% endraw %}+ + + +For the following lemmas, since maps are represented by functions, to +show two maps equal we will need to postulate extensionality. + +
+ + postulate + extensionality : ∀ {A : Set} {ρ ρ′ : TotalMap A} → (∀ x → ρ x ≡ ρ′ x) → ρ ≡ ρ′ + ++ +#### Exercise: 2 stars, optional (update-shadow) +If we update a map $$ρ$$ at a key $$x$$ with a value $$v$$ and then +update again with the same key $$x$$ and another value $$w$$, the +resulting map behaves the same (gives the same result when applied +to any key) as the simpler map obtained by performing just +the second update on $$ρ$$: + +
+ + postulate + update-shadow : ∀ {A} (ρ : TotalMap A) (x : Id) (v w : A) + → (ρ , x ↦ v , x ↦ w) ≡ (ρ , x ↦ w) + ++ +
+ + update-shadow′ : ∀ {A} (ρ : TotalMap A) (x : Id) (v w : A) + → ((ρ , x ↦ v) , x ↦ w) ≡ (ρ , x ↦ w) + update-shadow′ ρ x v w = extensionality lemma + where + lemma : ∀ y → ((ρ , x ↦ v) , x ↦ w) y ≡ (ρ , x ↦ w) y + lemma y with x ≟ y + ... | yes refl = refl + ... | no x≢y = update-neq ρ x v y x≢y + +
{% raw %} - + + postulate update-same : ∀ {A} (mρ : TotalMap A) (x y : Id) - → (updateρ m, x (m↦ ρ x))) y ≡ m yρ -{% endraw %}+ +
{% raw %} - + + update-same′ : ∀ {A} (mρ : TotalMap A) (x y : Id) - → (updateρ m, x (m↦ ρ x))) y ≡ m yρ update-same′ mρ x = extensionality ylemma withwhere + lemma x: ≟∀ y → (ρ , x ↦ ρ x) y ≡ ρ y - lemma y with x ≟ y + ... | yes x=y rewrite x=y = refl = refl - ... | no x≠y x≢y = refl -{% endraw %}+ +
{% raw %} - + + postulate update-permute : ∀ {A v1 v2} (mρ : TotalMap A) (x1x x2 y : Id) (v : A) (y : Id) (w : A) → x1x ≢ x2y → (updateρ (update, mx x2↦ v2)v x1, v1) y - ↦ w) ≡ (updateρ (update, m x1 v1) x2 v2) y ↦ w , x ↦ v) -{% endraw %}+ +
{% raw %} - + + update-permute′ : ∀ {A v1 v2} (mρ : TotalMap A) (x1x x2 y : Id) - → x1 ≢ x2 → (updatev (update: m x2 v2) x1 v1) y - ≡ (update (update m x1 v1) x2 v2) y - update-permute′ {A}) {v1} {v2} m x1 x2 (y x1≠x2: Id) (w : A) + → x ≢ y → (ρ , x ↦ v , y ↦ w) ≡ (ρ , y ↦ w , x ↦ v) + update-permute′ {A} ρ x v y w x≢y = extensionality lemma where + lemma : ∀ z → (ρ , x ↦ v , y ↦ w) z ≡ (ρ , y ↦ w , x ↦ v) z + lemma z with x1x ≟ yz | x2y ≟ y - ... | yes x1=y | yes x2=y = ⊥-elim (x1≠x2 (trans x1=y (sym x2=y))) - ... | no x1≠y | yes x2=y rewrite x2=y = update-eq′ m y - ... | yes x1=y | no x2≠y rewrite x1=y = sym (update-eq′ m y) - ... | no x1≠y | no x2≠y =z trans... | yes refl | yes refl = ⊥-elim (update-neqx≢y m x2 y x2≠yrefl) + ... (| no x≢z | yes refl = sym (update-eq′ ρ z w) + ... | yes refl | no y≢z = update-eq′ ρ z v + ... | no x≢z | no y≢z = trans (update-neq mρ x1x v z x≢z) (sym (update-neq ρ y x1≠yw z y≢z)) -{% endraw %}+ + + +And a slightly different version of the same proof. + +
+ + update-permute′′ : ∀ {A} (ρ : TotalMap A) (x : Id) (v : A) (y : Id) (w : A) (z : Id) + → x ≢ y → (ρ , x ↦ v , y ↦ w) z ≡ (ρ , y ↦ w , x ↦ v) z + update-permute′′ {A} ρ x v y w z x≢y with x ≟ z | y ≟ z + ... | yes x≡z | yes y≡z = ⊥-elim (x≢y (trans x≡z (sym y≡z))) + ... | no x≢z | yes y≡z rewrite y≡z = sym (update-eq′ ρ z w) + ... | yes x≡z | no y≢z rewrite x≡z = update-eq′ ρ z v + ... | no x≢z | no y≢z = trans (update-neq ρ x v z x≢z) (sym (update-neq ρ y w z y≢z)) + +
{% raw %} - + +PartialMap : Set → Set PartialMap A = TotalMap (Maybe A) -{% endraw %}-
{% raw %} - + +In this chapter, we develop the fundamental theory of the Simply Typed Lambda Calculus---in particular, the type safety @@ -299,326 +359,440 @@ is to identify the possible _canonical forms_ (i.e., well-typed closed values) belonging to each type. For $$bool$$, these are the boolean values $$true$$ and $$false$$. For arrow types, the canonical forms are lambda-abstractions. -+ +module PartialMap where -{% endraw %}-{% raw %} - empty + +-The simply typed lambda-calculus (STLC) is a tiny core -calculus embodying the key concept of _functional abstraction_, -which shows up in pretty much every real-world programming -language in some form (functions, procedures, methods, etc.). - -We will follow exactly the same pattern as in the previous chapter -when formalizing this calculus (syntax, small-step semantics, -typing rules) and its main properties (progress and preservation). -The new technical challenges arise from the mechanisms of -_variable binding_ and _substitution_. It which will take some -work to deal with these. - - -## Overview - -The STLC is built on some collection of _base types_: -booleans, numbers, strings, etc. The exact choice of base types -doesn't matter much---the construction of the language and its -theoretical properties work out the same no matter what we -choose---so for the sake of brevity let's take just $$bool$$ for -the moment. At the end of the chapter we'll see how to add more -base types, and in later chapters we'll enrich the pure STLC with -other useful constructs like pairs, records, subtyping, and -mutable state. - -Starting from boolean constants and conditionals, we add three -things: - - - variables - - function abstractions - - application - -This gives us the following collection of abstract syntax -constructors (written out first in informal BNF notation---we'll -formalize it below). - -$$ - \begin{array}{rll} - \text{Terms}\;s,t,u - ::= & x & \text{variable} \\ - \mid & \lambda x : A . t & \text{abstraction} \\ - \mid & s\;t & \text{application} \\ - \mid & true & \text{constant true} \\ - \mid & false & \text{constant false} \\ - \mid & \text{if }s\text{ then }t\text{ else }u & \text{conditional} - \end{array} -$$ - -In a lambda abstraction $$\lambda x : A . t$$, the variable $$x$$ is called the -_parameter_ to the function; the term $$t$$ is its _body_. The annotation $$:A$$ -specifies the type of arguments that the function can be applied to. - -Some examples: - - - The identity function for booleans: - - $$\lambda x:bool. x$$. - - The identity function for booleans, applied to the boolean $$true$$: - - $$(\lambda x:bool. x)\;true$$. - - The boolean "not" function: - - $$\lambda x:bool. \text{if }x\text{ then }false\text{ else }true$$. - - The constant function that takes every (boolean) argument to $$true$$: - - $$\lambda x:bool. true$$. - - A two-argument function that takes two booleans and returns the - first one: - - $$\lambda x:bool. \lambda y:bool. x$$. - - As in Agda, a two-argument function is really a - one-argument function whose body is also a one-argument function. - - A two-argument function that takes two booleans and returns the - first one, applied to the booleans $$false$$ and $$true$$: - - $$(\lambda x:bool. \lambda y:bool. x)\;false\;true$$. - - As in Agda, application associates to the left---i.e., this - expression is parsed as - - $$((\lambda x:bool. \lambda y:bool. x)\;false)\;true$$. - - - A higher-order function that takes a _function_ $$f$$ (from booleans - to booleans) as an argument, applies $$f$$ to $$true$$, and applies - $$f$$ again to the result: - - $$\lambda f:bool\rightarrow bool. f\;(f\;true)$$. - - - The same higher-order function, applied to the constantly $$false$$ - function: - - $$(\lambda f:bool\rightarrow bool. f\;(f\;true))\;(\lambda x:bool. false)$$. - -As the last several examples show, the STLC is a language of -_higher-order_ functions: we can write down functions that take -other functions as arguments and/or return other functions as -results. - -The STLC doesn't provide any primitive syntax for defining _named_ -functions---all functions are "anonymous." We'll see in chapter -`MoreStlc` that it is easy to add named functions to what we've -got---indeed, the fundamental naming and binding mechanisms are -exactly the same. - -The _types_ of the STLC include $$bool$$, which classifies the -boolean constants $$true$$ and $$false$$ as well as more complex -computations that yield booleans, plus _arrow types_ that classify -functions. - -$$ - \text{Types}\;A,B ::= bool \mid A \rightarrow B -$$ - -For example: - - - $$\lambda x:bool. false$$ has type $$bool\rightarrow bool$$; - - $$\lambda x:bool. x$$ has type $$bool\rightarrow bool$$; - - $$(\lambda x:bool. x)\;true$$ has type $$bool$$; - - $$\lambda x:bool. \lambda y:bool. x$$ has type - $$bool\rightarrow bool\rightarrow bool$$ - (i.e., $$bool\rightarrow (bool\rightarrow bool)$$) - - $$(\lambda x:bool. \lambda y:bool. x)\;false$$ has type $$bool\rightarrow bool$$ - - $$(\lambda x:bool. \lambda y:bool. x)\;false\;true$$ has type $$bool$$ ## Syntax -We begin by formalizing the syntax of the STLC. -Unfortunately, $$\rightarrow$$ is already used for Agda's function type, -so we will STLC's function type as `_⇒_`. +Syntax of types and terms. All source terms are labeled with $ᵀ$. ++ + ∅ : ∀ {A} → PartialMap A empty∅ = TotalMap.emptyTotalMap.always nothing -{% endraw %}-{% raw %} - update + ++As before, we define handy abbreviations for updating a map two, three, or four times. -+ + _,_↦_ : ∀ {A} (mρ : PartialMap A) (x : Id) (v : A) → PartialMap A updateρ m, x ↦ v = TotalMap.updateTotalMap._,_↦_ mρ x (just v) -{% endraw %}-We can now lift all of the basic lemmas about total maps to -partial maps. +{% raw %} - update-eq + + _,_↦_,_↦_ : ∀ {A v} (m→ : PartialMap A) (x→ : Id → A → Id → A → PartialMap A + ρ , x₁ ↦ v₁ , x₂ ↦ v₂ = (ρ , x₁ ↦ v₁), x₂ ↦ v₂ + + _,_↦_,_↦_,_↦_ : ∀ {A} → PartialMap A → Id → A → Id → A → Id → A → PartialMap A + ρ , x₁ ↦ v₁ , x₂ ↦ v₂ , x₃ ↦ v₃ = ((ρ , x₁ ↦ v₁), x₂ ↦ v₂), x₃ ↦ v₃ + + _,_↦_,_↦_,_↦_,_↦_ : ∀ {A} → PartialMap A → Id → A → Id → A → Id → A → Id → A → PartialMap A + ρ , x₁ ↦ v₁ , x₂ ↦ v₂ , x₃ ↦ v₃ , x₄ ↦ v₄ = (((ρ , x₁ ↦ v₁), x₂ ↦ v₂), x₃ ↦ v₃), x₄ ↦ v₄ + ++ +We now lift all of the basic lemmas about total maps to partial maps. + ++ + apply-∅ : ∀ {A} → (x : Id) → (∅ {A} x) ≡ nothing + apply-∅ x = TotalMap.apply-always nothing x + ++ ++ + update-eq : ∀ {A} (ρ : PartialMap A) (x : Id) (v : A) → (updateρ m, x ↦ v) x ≡ just v update-eq mρ x = TotalMap.update-eq m x -{% endraw %}- -{% raw %} - update-neq : ∀ {A v} (m= :TotalMap.update-eq PartialMapρ A) (x y(just : Idv) + ++ ++ + update-neq : ∀ {A} (ρ : PartialMap A) (x : Id) (v : A) (y : Id) → x ≢ y → (updateρ m, x ↦ v) y ≡ mρ y update-neq mρ x v y x≠yx≢y = TotalMap.update-neq mρ x y x≠y -{% endraw %}- -{% raw %} - update-shadow : ∀ {A v1 v2} (m : PartialMap A) (x y : Id) - → (update (update m x v1) x v2) y ≡ (update m x v2) y - update-shadow m x y = TotalMap.update-shadow m x y -{% endraw %}- -{% raw %} - update-same : ∀ {A v} (m : PartialMap A) (x y : Id) - → m x ≡ just v) y x≢y + ++ ++ + update-shadow : ∀ {A} (ρ : PartialMap A) (x : Id) (v w : A) + → (ρ , x ↦ v , x ↦ w) ≡ (ρ , x ↦ w) + update-shadow ρ x v w = TotalMap.update-shadow ρ x (just v) (just w) + ++ ++ + update-same : ∀ {A} (ρ : PartialMap A) (x : Id) (v : A) + → ρ x ≡ just v → (updateρ m, x ↦ v) y ≡ m yρ update-same mρ x yv mx=vρx≡v rewrite sym mx=vρx≡v = TotalMap.update-same mρ x y -{% endraw %}-{% raw %} - + +diff --git a/out/Stlc.md b/out/Stlc.md index d824d490..75e17557 100644 --- a/out/Stlc.md +++ b/out/Stlc.md @@ -4,5511 +4,4450 @@ layout : page permalink : /Stlc --- -+ + update-permute : ∀ {A v1 v2} (mρ : PartialMap A) (x1x x2 y : Id) + > (v : A) (y : Id) (w : A) → x1x ≢ x2y → (updateρ (update, mx x2↦ v2)v x1, v1) y - ↦ w) ≡ (updateρ (update, m x1 v1) x2 v2) y ↦ w , x ↦ v) update-permute mρ x1x x2v y x1≠x2w x≢y = TotalMap.update-permute mρ x1x x2(just v) y x1≠x2(just w) x≢y -{% endraw %}+ +-+{% raw %} - + +open import Maps using (Id; id; _≟_; PartialMap; moduleimport PartialMap) -openMaps import Data.Empty using (⊥Id; ⊥-elimid; _≟_; PartialMap; module PartialMap) open importPartialMap Data.Maybe using (∅; (Maybe; just; nothing_,_↦_) open import Data.NatData.String using (ℕ; suc; zero; _+_String) open import Data.Empty using import Data.Product using (∃⊥; ∄; _,_⊥-elim) open import FunctionData.Maybe using (_∘_Maybe; _$_just; nothing) +open import -open importData.Nat Relation.Nullary using (ℕ; suc; zero; _+_) (Dec; yes; no) open import Relation.Binary.PropositionalEqualityData.Bool using (_≡_Bool; _≢_true; reflfalse; if_then_else_) +open import Relation.Nullary using (Dec; yes; no) +open import Relation.Nullary.Decidable using (⌊_⌋) +open import Relation.Binary.PropositionalEquality using (_≡_; _≢_; refl) +-- open import Relation.Binary.Core using (Rel) +-- open import Data.Product using (∃; ∄; _,_) +-- open import Function using (_∘_; _$_) -{% endraw %}--### Types - -diff --git a/out/StlcProp.md b/out/StlcProp.md index 84676601..fce8d296 100644 --- a/out/StlcProp.md +++ b/out/StlcProp.md @@ -4,287 +4,347 @@ layout : page permalink : /StlcProp --- -{% raw %} -data Type : Set where - bool : Type - _⇒_ : Type → Type → Type - -infixr 5100 _⇒_ -{% endraw %}- - -### Terms - -{% raw %} -data Term : Set where - var : Id → Term - app : Term → Term → Term - abs : Id → Type → Term → Term - true : Term - false : Term - if_then_else_ : Term → Term → Term → Term -{% endraw %}- --- -Note that an abstraction $$\lambda x:A.t$$ (formally, `abs x A t`) is -always annotated with the type $$A$$ of its parameter, in contrast -to Agda (and other functional languages like ML, Haskell, etc.), -which use _type inference_ to fill in missing annotations. We're -not considering type inference here. - -We introduce $$x, y, z$$ as names for variables. The pragmas ensure -that $$id 0, id 1, id 2$$ display as $$x, y, z$$. - -{% raw %} -infixr 8 if_then_else_ -{% endraw %}-{% raw %} -x = id 0 y = id 1 -z = id 2 - -{-# DISPLAY id zero = x #-} -{-# DISPLAY id (suc zero) = y #-} -{-# DISPLAY id (suc (suc zero)) = z #-} -{% endraw %}- -Some examples... - -$$\text{idB} = \lambda x:bool. x$$. - -{% raw %} -idB = (abs x bool (var x)) -{% endraw %}- -$$\text{idBB} = \lambda x:bool\rightarrow bool. x$$. - -{% raw %} -idBB = (abs x (bool ⇒ bool) (var x)) -{% endraw %}- -$$\text{idBBBB} = \lambda x:(bool\rightarrow bool)\rightarrow (bool\rightarrow bool). x$$. - -{% raw %} -idBBBB = (abs x ((bool ⇒ bool) ⇒ (bool ⇒ bool)) (var x)) -{% endraw %}- -$$\text{k} = \lambda x:bool. \lambda y:bool. x$$. - -{% raw %} -k = (abs x bool (abs y bool (var x))) -{% endraw %}- -$$\text{notB} = \lambda x:bool. \text{if }x\text{ then }false\text{ else }true$$. - -{% raw %} -notB = (abs x bool (if (var x) then false else true)) -{% endraw %}- --- - -## Operational Semantics - -To define the small-step semantics of STLC terms, we begin, -as always, by defining the set of values. Next, we define the -critical notions of _free variables_ and _substitution_, which are -used in the reduction rule for application expressions. And -finally we give the small-step relation itself. - -### Values - -To define the values of the STLC, we have a few cases to consider. - -First, for the boolean part of the language, the situation is -clear: $$true$$ and $$false$$ are the only values. An $$\text{if}$$ -expression is never a value. - -Second, an application is clearly not a value: It represents a -function being invoked on some argument, which clearly still has -work left to do. - -Third, for abstractions, we have a choice: - - - We can say that $$\lambda x:A. t$$ is a value only when $$t$$ is a - value---i.e., only if the function's body has been - reduced (as much as it can be without knowing what argument it - is going to be applied to). - - - Or we can say that $$\lambda x:A. t$$ is always a value, no matter - whether $$t$$ is one or not---in other words, we can say that - reduction stops at abstractions. - -Agda makes the first choice---for example, - -{% raw %} -{-# DISPLAY abs 0 bool (var 0) = idB #-} -{-# DISPLAY abs 0 (bool ⇒ bool) (var 0) = idBB #-} -{-# DISPLAY abs 0 ((bool ⇒ bool) ⇒ (bool ⇒ bool)) (var 0) = idBBBB #-} -{-# DISPLAY abs 0 bool (abs y bool (var 0)) = k #-} -{-# DISPLAY abs 0 bool (if (var 0) then false else true) = notB #-} -{% endraw %}-{% raw %} -test_normalizeUnderLambda : (λ (x : ℕ) → 3 + 4) ≡ (λ (x : ℕ) → 7) -test_normalizeUnderLambda = refl -{% endraw %}- -Most real-world functional programming languages make the second -choice---reduction of a function's body only begins when the -function is actually applied to an argument. We also make the -second choice here. - -{% raw %} -data Value : Term → Set where - abs : ∀ {x A t} - → Value (abs x A t) - true : Value true - false : Value false -{% endraw %}- -Finally, we must consider what constitutes a _complete_ program. - -Intuitively, a "complete program" must not refer to any undefined -variables. We'll see shortly how to define the _free_ variables -in a STLC term. A complete program is _closed_---that is, it -contains no free variables. - -Having made the choice not to reduce under abstractions, we don't -need to worry about whether variables are values, since we'll -always be reducing programs "from the outside in," and that means -the small-step relation will always be working with closed terms. - - -### Substitution - -Now we come to the heart of the STLC: the operation of -substituting one term for a variable in another term. This -operation is used below to define the operational semantics of -function application, where we will need to substitute the -argument term for the function parameter in the function's body. -For example, we reduce - -$$(\lambda x:bool. \text{if }x\text{ then }true\text{ else }x)\;false$$ - -to - -$$\text{if }false\text{ then }true\text{ else }false$$ - -by substituting $$false$$ for the parameter $$x$$ in the body of the -function. - -In general, we need to be able to substitute some given term $$s$$ -for occurrences of some variable $$x$$ in another term $$t$$. In -informal discussions, this is usually written $$[x:=s]t$$ and -pronounced "substitute $$x$$ with $$s$$ in $$t$$." - -Here are some examples: - - - $$[x:=true](\text{if }x\text{ then }x\text{ else }false)$$ - yields $$\text{if }true\text{ then }true\text{ else }false$$ - - - $$[x:=true]x$$ - yields $$true$$ - - - $$[x:=true](\text{if }x\text{ then }x\text{ else }y)$$ - yields $$\text{if }true\text{ then }true\text{ else }y$$ - - - $$[x:=true]y$$ - yields $$y$$ - - - $$[x:=true]false$$ - yields $$false$$ (vacuous substitution) - - - $$[x:=true](\lambda y:bool. \text{if }y\text{ then }x\text{ else }false)$$ - yields $$\lambda y:bool. \text{if }y\text{ then }true\text{ else }false$$ - - - $$[x:=true](\lambda y:bool. x)$$ - yields $$\lambda y:bool. true$$ - - - $$[x:=true](\lambda y:bool. y)$$ - yields $$\lambda y:bool. y$$ - - - $$[x:=true](\lambda x:bool. x)$$ - yields $$\lambda x:bool. x$$ - -The last example is very important: substituting $$x$$ with $$true$$ in -$$\lambda x:bool. x$$ does _not_ yield $$\lambda x:bool. true$$! The reason for -this is that the $$x$$ in the body of $$\lambda x:bool. x$$ is _bound_ by the -abstraction: it is a new, local name that just happens to be -spelled the same as some global name $$x$$. - -Here is the definition, informally... - -$$ - \begin{aligned} - &[x:=s]x &&= s \\ - &[x:=s]y &&= y \;\{\text{if }x\neq y\} \\ - &[x:=s](\lambda x:A. t) &&= \lambda x:A. t \\ - &[x:=s](\lambda y:A. t) &&= \lambda y:A. [x:=s]t \;\{\text{if }x\neq y\} \\ - &[x:=s](t1\;t2) &&= ([x:=s]t1) ([x:=s]t2) \\ - &[x:=s]true &&= true \\ - &[x:=s]false &&= false \\ - &[x:=s](\text{if }t1\text{ then }t2\text{ else }t3) &&= - \text{if }[x:=s]t1\text{ then }[x:=s]t2\text{ else }[x:=s]t3 - \end{aligned} -$$ - -... and formally: - -{% raw %} -[_:=_]_ : Id -> Term -> Term -> Term -[ x := v ] (var y) with x ≟ y -... | yes x=y = v -... | no x≠y = var y -[ x := v ] (app s t) = app ([ x := v ] s) ([ x := v ] t) -[ x := v ] (abs y A t) with x ≟ y -... | yes x=y = abs y A t -... | no x≠y = abs y A ([ x := v ] t) -[ x := v ] true = true -[ x := v ] false = false -[ x := v ] (if s then t else u) = - if [ x := v ] s then [ x := v ] t else [ x := v ] u -{% endraw %}- - - -_Technical note_: Substitution becomes trickier to define if we -consider the case where $$s$$, the term being substituted for a -variable in some other term, may itself contain free variables. -Since we are only interested here in defining the small-step relation -on closed terms (i.e., terms like $$\lambda x:bool. x$$ that include -binders for all of the variables they mention), we can avoid this -extra complexity here, but it must be dealt with when formalizing -richer languages. - - -#### Exercise: 3 stars (subst-correct) -The definition that we gave above defines substitution as a _function_. -Suppose, instead, we wanted to define substitution as an inductive _relation_. -We've begun the definition by providing the `data` header and -one of the constructors; your job is to fill in the rest of the constructors -and prove that the relation you've defined coincides with the function given -above. -{% raw %} -data [_:=_]_==>_ (x : Id) (s : Term) : Term -> Term -> Set where - var1 : [ x := s ] (var x) ==> s - {- FILL IN HERE -} -{% endraw %}- -{% raw %} -postulate - subst-correct : ∀ s x t t' - → [ x := s ] t ≡ t' - → [ x := s ] t ==> t' -{% endraw %}- -### Reduction - -The small-step reduction relation for STLC now follows the -same pattern as the ones we have seen before. Intuitively, to -reduce a function application, we first reduce its left-hand -side (the function) until it becomes an abstraction; then we -reduce its right-hand side (the argument) until it is also a -value; and finally we substitute the argument for the bound -variable in the body of the abstraction. This last rule, written -informally as - -$$ -(\lambda x : A . t) v \Longrightarrow [x:=v]t -$$ - -is traditionally called "beta-reduction". - -$$ -\begin{array}{cl} - \frac{value(v)}{(\lambda x : A . t) v \Longrightarrow [x:=v]t}&(red)\\\\ - \frac{s \Longrightarrow s'}{s\;t \Longrightarrow s'\;t}&(app1)\\\\ - \frac{value(v)\quad t \Longrightarrow t'}{v\;t \Longrightarrow v\;t'}&(app2) -\end{array} -$$ - -... plus the usual rules for booleans: - -$$ -\begin{array}{cl} - \frac{}{(\text{if }true\text{ then }t_1\text{ else }t_2) \Longrightarrow t_1}&(iftrue)\\\\ - \frac{}{(\text{if }false\text{ then }t_1\text{ else }t_2) \Longrightarrow t_2}&(iffalse)\\\\ - \frac{s \Longrightarrow s'}{\text{if }s\text{ then }t_1\text{ else }t_2 - \Longrightarrow \text{if }s\text{ then }t_1\text{ else }t_2}&(if) -\end{array} -$$ - -Formally: - -{% raw %} -data _==>_ : Term → Term → Set where - red : ∀ {x A s t} - → Value t - → app (abs x A s) t ==> [ x := t ] s - app1 : ∀ {s s' t} - → s ==> s' - → app s t ==> app s' t - app2 : ∀ {s t t'} - → Value s - → t ==> t' - → app s t ==> app s t' - if : ∀ {s s' t u} - → s ==> s' - → if s then t else u ==> if s' then t else u - iftrue : ∀ {s t} - → if true then s else t ==> s - iffalse : ∀ {s t} - → if false then s else t ==> t -{% endraw %}- - - -{% raw %} -data Multi (R : Term → Term → Set) : Term → Term → Set where - refl : ∀ {x} -> Multi R x x - step : ∀ {x y z} -> R x y -> Multi R y z -> Multi R x z -{% endraw %}- -{% raw %} -_==>*_ : Term → Term → Set -_==>*_ = Multi _==>_ -{% endraw %}- - - -### Examples - -Example: - -$$((\lambda x:bool\rightarrow bool. x) (\lambda x:bool. x)) \Longrightarrow^* (\lambda x:bool. x)$$. - -{% raw %} -step-example1 : (app idBB idB) ==>* idB -step-example1 = step (red abs) - $ refl -{% endraw %}- -Example: - -$$(\lambda x:bool\rightarrow bool. x) \;((\lambda x:bool\rightarrow bool. x)\;(\lambda x:bool. x))) \Longrightarrow^* (\lambda x:bool. x)$$. - -{% raw %} -step-example2 : (app idBB (app idBB idB)) ==>* idB -step-example2 = step (app2 abs (red abs)) - $ step (red abs) - $ refl -{% endraw %}- -Example: - -$$((\lambda x:bool\rightarrow bool. x)\;(\lambda x:bool. \text{if }x\text{ then }false\text{ else }true))\;true\Longrightarrow^* false$$. - -{% raw %} -step-example3 : (app (app idBB notB) true) ==>* false -step-example3 = step (app1 (red abs)) - $ step (red true) - $ step iftrue - $ refl -{% endraw %}- -Example: - -$$((\lambda x:bool\rightarrow bool. x)\;((\lambda x:bool. \text{if }x\text{ then }false\text{ else }true)\;true))\Longrightarrow^* false$$. - -{% raw %} -step-example4 : (app idBB (app notB true)) ==>* false -step-example4 = step (app2 abs (red true)) - $ step (app2 abs iftrue) - $ step (red false) - $ refl -{% endraw %}- -#### Exercise: 2 stars (step-example5) - -{% raw %} -postulate - step-example5 : (app (app idBBBB idBB) idB) ==>* idB -{% endraw %}- -## Typing - -Next we consider the typing relation of the STLC. - -### Contexts - -_Question_: What is the type of the term "$$x\;y$$"? - -_Answer_: It depends on the types of $$x$$ and $$y$$! - -I.e., in order to assign a type to a term, we need to know -what assumptions we should make about the types of its free -variables. - -This leads us to a three-place _typing judgment_, informally -written $$\Gamma\vdash t : A$$, where $$\Gamma$$ is a -"typing context"---a mapping from variables to their types. - -Informally, we'll write $$\Gamma , x:A$$ for "extend the partial function -$$\Gamma$$ to also map $$x$$ to $$A$$." Formally, we use the function `_,_∶_` -(or "update") to add a binding to a context. - -{% raw %} -Ctxt : Set -Ctxt = PartialMap Type - -∅ : Ctxt -∅ = PartialMap.empty - -_,_∶_ : Ctxt -> Id -> Type -> Ctxt -_,_∶_ = PartialMap.update -{% endraw %}- - - -### Typing Relation - -$$ - \begin{array}{cl} - \frac{\Gamma\;x = A}{\Gamma\vdash{x:A}}&(var)\\\\ - \frac{\Gamma,x:A\vdash t:B}{\Gamma\vdash (\lambda x:A.t) : A\rightarrow B}&(abs)\\\\ - \frac{\Gamma\vdash s:A\rightarrow B\quad\Gamma\vdash t:A}{\Gamma\vdash (s\;t) : B}&(app)\\\\ - \frac{}{\Gamma\vdash true : bool}&(true)\\\\ - \frac{}{\Gamma\vdash false : bool}&(true)\\\\ - \frac{\Gamma\vdash s:bool \quad \Gamma\vdash t1:A \quad \Gamma\vdash t2:A}{\Gamma\vdash\text{if }s\text{ then }t1\text{ else }t2 : A}&(if) - \end{array} -$$ - -We can read the three-place relation $$\Gamma\vdash (t : A)$$ as: -"to the term $$t$$ we can assign the type $$A$$ using as types for -the free variables of $$t$$ the ones specified in the context -$$\Gamma$$." - -{% raw %} -data _⊢_∶_ : Ctxt -> Term -> Type ->: Set where var 𝔹 : ∀ {Γ} x {A} - → Γ x ≡ just A - → Γ ⊢ var x ∶ AType abs _⇒_ : ∀Type {Γ} {x} {A} {B} {s} - → ΓType , x ∶ A ⊢ s ∶ B - → ΓType + +data ⊢Term abs: xSet A s ∶ A ⇒ Bwhere app varᵀ : ∀Id {Γ} {A} {B} {s} {t} - → Γ ⊢ s ∶ A ⇒ B - → Γ ⊢ t ∶ A - → Γ ⊢ app s t ∶ BTerm true λᵀ_∈_⇒_ : ∀Id {Γ} - → ΓType ⊢→ trueTerm → Term + _·ᵀ_ : Term → Term → Term + trueᵀ : Term + falseᵀ : Term + ifᵀ_then_else_ : Term → Term → Term → Term + ++ +Some examples. ++ +f x y : Id +f ∶= id bool - false : ∀ {Γ} - → Γ ⊢ false ∶ bool - if_then_else_ : ∀ {Γ} {s} {t} {u} {A} - → Γ ⊢ s ∶ bool - → Γ ⊢ t ∶ A - → Γ ⊢ u ∶ A - → Γ ⊢ if s then t else u ∶ A -{% endraw %}- - - - -### Examples - -{% raw %} -typing-example1 : ∅ ⊢ idB ∶ bool ⇒ bool"f" typing-example1 x = id abs (var x refl) -{% endraw %}- -Another example: - -$$\varnothing\vdash \lambda x:A. \lambda y:A\rightarrow A. y\;(y\;x) : A\rightarrow (A\rightarrow A)\rightarrow A$$. - -{% raw %} -typing-example2 : ∅ ⊢ - (abs x bool - (abs y (bool ⇒ bool) - (app (var y) - (app (var y) (var x))))) - ∶ (bool ⇒ (bool ⇒ bool) ⇒ bool)"x" typing-example2 y = id "y" - (abs - (abs - (app (var y refl) - (app (var y refl) (var x refl) )))) -{% endraw %}-#### Exercise: 2 stars (typing-example3) -Formally prove the following typing derivation holds: - -$$\exists A, \varnothing\vdash \lambda x:bool\rightarrow B. \lambda y:bool\rightarrow bool. \lambda z:bool. y\;(x\;z) : A$$. - -{% raw %} -postulate - typing-example3I[𝔹] I[𝔹⇒𝔹] K[𝔹][𝔹] not[𝔹] : ∃ λ A → ∅ ⊢ - Term +I[𝔹] = (absλᵀ x (bool∈ 𝔹 ⇒ bool)(varᵀ x)) - I[𝔹⇒𝔹] = (absλᵀ yf ∈ (bool𝔹 ⇒ bool𝔹) - ⇒ (absλᵀ z bool - (app (var y) (app (var x ∈ 𝔹 ⇒ ((varᵀ f) ·ᵀ (varvarᵀ z)))))) ∶ A -{% endraw %}- -We can also show that terms are _not_ typable. For example, let's -formally check that there is no typing derivation assigning a type -to the term $$\lambda x:bool. \lambda y:bool. x\;y$$---i.e., - - -$$\nexists A, \varnothing\vdash \lambda x:bool. \lambda y:bool. x\;y : A$$. - -{% raw %} -postulate - typing-nonexample1 : ∄ λ A → ∅ ⊢ - (abs x bool - (abs y bool - (app (var x) (var y)))) ∶ A -{% endraw %}- -#### Exercise: 3 stars, optional (typing-nonexample2) -Another nonexample: - -$$\nexists A, \exists B, \varnothing\vdash \lambda x:A. x\;x : B$$. - -{% raw %} -postulate - typing-nonexample2 : ∄ λ A → ∃ λ B → ∅ ⊢ - K[𝔹][𝔹] = (absλᵀ x B∈ 𝔹 ⇒ (appλᵀ y ∈ 𝔹 ⇒ (varvarᵀ x) (var x))) +not[𝔹] = (λᵀ ∶x ∈ 𝔹 ⇒ (ifᵀ (varᵀ x) then falseᵀ else trueᵀ)) + ++ +## Values + ++ +data value : Term → Set where + value-λᵀ : ∀ {x A N} → value (λᵀ x ∈ A ⇒ N) + value-trueᵀ : value (trueᵀ) + value-falseᵀ : value (falseᵀ) + ++ +## Substitution + ++ +_[_:=_] : Term → Id → Term → Term +(varᵀ x′) [ x := V ] with x ≟ x′ +... | yes _ = V +... | no _ = varᵀ x′ +(λᵀ x′ ∈ A′ ⇒ N′) [ x := V ] with x ≟ x′ +... | yes _ = λᵀ x′ ∈ A′ ⇒ N′ +... | no _ = λᵀ x′ ∈ A′ ⇒ (N′ [ x := V ]) +(L′ ·ᵀ M′) [ x := V ] = (L′ [ x := V ]) ·ᵀ (M′ [ x := V ]) +(trueᵀ) [ x := V ] = trueᵀ +(falseᵀ) [ x := V ] = falseᵀ +(ifᵀ L′ then M′ else N′) [ x := V ] = ifᵀ (L′ [ x := V ]) then (M′ [ x := V ]) else (N′ [ x := V ]) + ++ +## Reduction rules + ++ +data _⟹_ : Term → Term → Set where + β⇒ : ∀ {x A N V} → value V → + ((λᵀ x ∈ A ⇒ N) ·ᵀ V) ⟹ (N [ x := V ]) + γ⇒₁ : ∀ {L L' M} → + L ⟹ L' → + (L ·ᵀ M) ⟹ (L' ·ᵀ M) + γ⇒₂ : ∀ {V M M'} → + value V → + M ⟹ M' → + (V ·ᵀ M) ⟹ (V ·ᵀ M') + β𝔹₁ : ∀ {M N} → + (ifᵀ trueᵀ then M else N) ⟹ M + β𝔹₂ : ∀ {M N} → + (ifᵀ falseᵀ then M else N) ⟹ N + γ𝔹 : ∀ {L L' M N} → + L ⟹ L' → + (ifᵀ L then M else N) ⟹ (ifᵀ L' then M else N) + ++ +## Reflexive and transitive closure of a relation + ++ +Rel : Set → Set₁ +Rel A = A → A → Set + +infixl 100 _>>_ + +data _* {A : Set} (R : Rel A) : Rel A where + ⟨⟩ : ∀ {x : A} → (R *) x x + ⟨_⟩ : ∀ {x y : A} → R x y → (R *) x y + _>>_ : ∀ {x y z : A} → (R *) x y → (R *) y z → (R *) x z + ++ ++ +_⟹*_ : Term → Term → Set +_⟹*_ = (_⟹_) * + ++ +Example evaluation. + ++ +example₀ : (not[𝔹] ·ᵀ trueᵀ) ⟹* falseᵀ +example₀ = ⟨ step₀ ⟩ >> ⟨ step₁ ⟩ + where + M₀ M₁ M₂ : Term + M₀ = (not[𝔹] ·ᵀ trueᵀ) + M₁ = (ifᵀ trueᵀ then falseᵀ else trueᵀ) + M₂ = falseᵀ + step₀ : M₀ ⟹ M₁ + step₀ = β⇒ value-trueᵀ + step₁ : M₁ ⟹ M₂ + step₁ = β𝔹₁ + +example₁ : (I[𝔹⇒𝔹] ·ᵀ I[𝔹] ·ᵀ (not[𝔹] ·ᵀ falseᵀ)) ⟹* trueᵀ +example₁ = ⟨ step₀ ⟩ >> ⟨ step₁ ⟩ >> ⟨ step₂ ⟩ >> ⟨ step₃ ⟩ >> ⟨ step₄ ⟩ + where + M₀ M₁ M₂ M₃ M₄ M₅ : Term + M₀ = (I[𝔹⇒𝔹] ·ᵀ I[𝔹] ·ᵀ (not[𝔹] ·ᵀ falseᵀ)) + M₁ = ((λᵀ x ∈ 𝔹 ⇒ (I[𝔹] ·ᵀ varᵀ x)) ·ᵀ (not[𝔹] ·ᵀ falseᵀ)) + M₂ = ((λᵀ x ∈ 𝔹 ⇒ (I[𝔹] ·ᵀ varᵀ x)) ·ᵀ (ifᵀ falseᵀ then falseᵀ else trueᵀ)) + M₃ = ((λᵀ x ∈ 𝔹 ⇒ (I[𝔹] ·ᵀ varᵀ x)) ·ᵀ trueᵀ) + M₄ = I[𝔹] ·ᵀ trueᵀ + M₅ = trueᵀ + step₀ : M₀ ⟹ M₁ + step₀ = γ⇒₁ (β⇒ value-λᵀ) + step₁ : M₁ ⟹ M₂ + step₁ = γ⇒₂ value-λᵀ (β⇒ value-falseᵀ) + step₂ : M₂ ⟹ M₃ + step₂ = γ⇒₂ value-λᵀ β𝔹₂ + step₃ : M₃ ⟹ M₄ + step₃ = β⇒ value-trueᵀ + step₄ : M₄ ⟹ M₅ + step₄ = β⇒ value-trueᵀ + ++ +## Type rules + ++ +Context : Set +Context = PartialMap Type + +data _⊢_∈_ : Context → Term → Type → Set where + Ax : ∀ {Γ x A} → + Γ x ≡ just A → + Γ ⊢ varᵀ x ∈ A + ⇒-I : ∀ {Γ x N A B} → + (Γ , x ↦ A) ⊢ N ∈ B → + Γ ⊢ (λᵀ x ∈ A ⇒ N) ∈ (A ⇒ B) + ⇒-E : ∀ {Γ L M A B} → + Γ ⊢ L ∈ (A ⇒ B) → + Γ ⊢ M ∈ A → + Γ ⊢ L ·ᵀ M ∈ B + 𝔹-I₁ : ∀ {Γ} → + Γ ⊢ trueᵀ ∈ 𝔹 + 𝔹-I₂ : ∀ {Γ} → + Γ ⊢ falseᵀ ∈ 𝔹 + 𝔹-E : ∀ {Γ L M N A} → + Γ ⊢ L ∈ 𝔹 → + Γ ⊢ M ∈ A → + Γ ⊢ N ∈ A → + Γ ⊢ (ifᵀ L then M else N) ∈ A -{% endraw %}+ +-+ +{% raw %} - + +open import Function using (_∘_) open import Data.Empty using (⊥; ⊥-elim) +open import Data.Bool using (⊥; ⊥-elim) -open import(Bool; Data.Maybe using (Maybetrue; false) +open import just;Data.Maybe nothing) -open import Data.Product using (∃Maybe; ∃₂just; _,_; ,_nothing) open import Data.Product using (∃; Data.Sum∃₂; _,_; using,_) +open import Data.Sum using (_⊎_; inj₁; inj₂) open import Relation.Nullary using (¬_; Dec; yes; no) - open import Relation.Binary.PropositionalEquality using (_≡_¬_; _≢_Dec; reflyes; no) open import Relation.Binary.PropositionalEquality using (_≡_; _≢_; refl; trans; sym) +open import Maps open Maps.PartialMap +open import Stlc -{% endraw %}-
{% raw %} -CanonicalForms + +data : Term → Type → Set -CanonicalForms t bool = t ≡ truecanonical_for_ ⊎: tTerm ≡ false -CanonicalForms t (A ⇒ B) = ∃₂ λ x t′ → tType ≡→ Set where + canonical-λᵀ : ∀ {x A N B} abs→ canonical (λᵀ x ∈ A t′ - -canonicalForms⇒ N) : ∀for {t(A A}⇒ B→) ∅ ⊢ t ∶ A → Value + canonical-trueᵀ t: →canonical CanonicalFormstrueᵀ tfor A𝔹 -canonicalForms (abscanonical-falseᵀ t′): abs =canonical _falseᵀ ,for _ , refl𝔹 + canonicalForms true true = inj₁ refl-- canonical_for_ : Term → Type → Set canonicalForms-- canonical L for 𝔹 = L ≡ trueᵀ ⊎ L ≡ falseᵀ +-- canonical L for (A ⇒ B) = ∃₂ λ x N → L ≡ λᵀ x ∈ A ⇒ N + +canonicalFormsLemma false: ∀ {L A} → ∅ ⊢ L ∈ A → value L → canonical L for A +canonicalFormsLemma (Ax ⊢x) () +canonicalFormsLemma (⇒-I ⊢N) value-λᵀ = canonical-λᵀ -- _ , _ , refl +canonicalFormsLemma (⇒-E ⊢L ⊢M) () +canonicalFormsLemma 𝔹-I₁ value-trueᵀ = canonical-trueᵀ false = inj₂ refl-- inj₁ refl -{% endraw %}+canonicalFormsLemma 𝔹-I₂ value-falseᵀ = canonical-falseᵀ -- inj₂ refl +canonicalFormsLemma (𝔹-E ⊢L ⊢M ⊢N) () + + ## Progress @@ -629,99 +803,101 @@ straightforward extension of the progress proof we saw in the [Stlc]({{ "Stlc" | relative_url }}) chapter. We'll give the proof in English first, then the formal version. -
{% raw %} - + +progress : ∀ {tM A} → ∅ ⊢ tM ∶∈ A → Valuevalue tM ⊎ ∃ λ t′N → tM ==>⟹ t′N -{% endraw %}+ + _Proof_: By induction on the derivation of $$\vdash t : A$$. @@ -761,675 +937,830 @@ _Proof_: By induction on the derivation of $$\vdash t : A$$. - Otherwise, $$t_1$$ takes a step, and therefore so does $$t$$ (by `if`). -
{% raw %} - + +progress (var x ()) -progress true = inj₁ true -progress false = inj₁ false -progress (abs t∶A) = inj₁ abs -progress (app t₁∶A⇒B t₂∶B) - with progress t₁∶A⇒B -... | inj₂ (_ , t₁⇒t₁′) = inj₂ (_ , app1 t₁⇒t₁′) -... | inj₁ v₁ - with progress t₂∶B -... | inj₂ (_ , t₂⇒t₂′) = inj₂ (_ , app2 v₁ t₂⇒t₂′) -... | inj₁ v₂ - with canonicalForms t₁∶A⇒B v₁ -... | (_ , _ , refl) = inj₂ (_ , red v₂) -progress (if t₁∶bool then t₂∶A else t₃∶A) - with progress t₁∶bool -... |(Ax inj₂()) +progress (_ , t₁⇒t₁′) =(⇒-I inj₂ (_ , if t₁⇒t₁′⊢N) -... | inj₁ v₁ - with canonicalForms t₁∶bool v₁ -... | inj₁ refl = inj₁ value-λᵀ +progress (⇒-E {Γ} {L} {M} {A} {B} ⊢L ⊢M) with progress ⊢L +... | inj₂ (L′ (_ , iftrueL⟹L′) -... |= inj₂ (L′ inj₂·ᵀ M refl = inj₂ (_ , iffalseγ⇒₁ 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 , β𝔹₂) -{% endraw %}+ + #### Exercise: 3 stars, optional (progress_from_term_ind) Show that progress can also be proved by induction on terms instead of induction on typing derivations. -
{% raw %} - + +postulate progress′ : ∀ {tM A} → ∅ ⊢ tM ∶∈ A → Valuevalue tM ⊎ ∃ λ t′N → tM ==>⟹ t′N -{% endraw %}+ + ## Preservation @@ -1475,633 +1806,693 @@ order... ### Free Occurrences -A variable $$x$$ _appears free in_ a term _t_ if $$t$$ contains some -occurrence of $$x$$ that is not under an abstraction labeled $$x$$. +A variable $$x$$ _appears free in_ a term $$M$$ if $$M$$ contains some +occurrence of $$x$$ that is not under an abstraction over $$x$$. For example: - - $$y$$ appears free, but $$x$$ does not, in $$\lambda x : A\to B. x\;y$$ - - both $$x$$ and $$y$$ appear free in $$(\lambda x:A\to B. x\;y) x$$ - - no variables appear free in $$\lambda x:A\to B. \lambda y:A. x\;y$$ + - $$y$$ appears free, but $$x$$ does not, in $$λᵀ x ∈ (A ⇒ B) ⇒ x ·ᵀ y$$ + - both $$x$$ and $$y$$ appear free in $$(λᵀ x ∈ (A ⇒ B) ⇒ x ·ᵀ y) ·ᵀ x$$ + - no variables appear free in $$λᵀ x ∈ (A ⇒ B) ⇒ (λᵀ y ∈ A ⇒ x ·ᵀ y)$$ Formally: -
{% raw %} - + +data _FreeIn_ (x : Id) :→ Term → Set where varfree-varᵀ : x FreeIn var x - abs : ∀ {y A tx} → x FreeIn (varᵀ x) + free-λᵀ : ∀ {x y A N} → y ≢ x → x FreeIn tN → x FreeIn abs(λᵀ y ∈ A t - app1⇒ : ∀ {t₁ t₂} → x FreeIn t₁ → x FreeIn app t₁ t₂ - app2 : ∀ {t₁ t₂} → x FreeIn t₂ → x FreeIn app t₁ t₂ - if1 : ∀ {t₁ t₂ t₃} → x FreeIn t₁ → x FreeIn (if t₁ then t₂ else t₃N) if2 free-·ᵀ₁ : ∀ {t₁x t₂L t₃M} → x FreeIn t₂L → x FreeIn (ifL t₁·ᵀ then t₂ else t₃M) if3 free-·ᵀ₂ : ∀ {t₁x t₂L t₃M} → x FreeIn t₃M → x FreeIn (ifL t₁·ᵀ M) + free-ifᵀ₁ : ∀ {x L M N} → x FreeIn L → x FreeIn (ifᵀ L then t₂M else t₃N) + free-ifᵀ₂ : ∀ {x L M N} → x FreeIn M → x FreeIn (ifᵀ L then M else N) + free-ifᵀ₃ : ∀ {x L M N} → x FreeIn N → x FreeIn (ifᵀ L then M else N) -{% endraw %}+ + A term in which no variables appear free is said to be _closed_. -
{% raw %} -Closed + +closed : Term → Set Closedclosed tM = ∀ {x} → ¬ (x FreeIn tM) -{% endraw %}+ + #### Exercise: 1 star (free-in) If the definition of `_FreeIn_` is not crystal clear to @@ -2114,1600 +2505,1139 @@ are really the crux of the lambda-calculus.) ### Substitution To prove that substitution preserves typing, we first need a technical lemma connecting free variables and typing contexts: If -a variable $$x$$ appears free in a term $$t$$, and if we know $$t$$ is -well typed in context $$\Gamma$$, then it must be the case that -$$\Gamma$$ assigns a type to $$x$$. +a variable $$x$$ appears free in a term $$M$$, and if we know $$M$$ is +well typed in context $$Γ$$, then it must be the case that +$$Γ$$ assigns a type to $$x$$. -
{% raw %} -freeInCtxt + +freeLemma : ∀ {x tM A Γ} → x FreeIn tM → Γ ⊢ tM ∶∈ A → ∃ λ B → Γ x ≡ just B -{% endraw %}+ + _Proof_: We show, by induction on the proof that $$x$$ appears - free in $$t$$, that, for all contexts $$\Gamma$$, if $$t$$ is well - typed under $$\Gamma$$, then $$\Gamma$$ assigns some type to $$x$$. + free in $$P$$, that, for all contexts $$Γ$$, if $$P$$ is well + typed under $$Γ$$, then $$Γ$$ assigns some type to $$x$$. - - If the last rule used was `var`, then $$t = x$$, and from - the assumption that $$t$$ is well typed under $$\Gamma$$ we have - immediately that $$\Gamma$$ assigns a type to $$x$$. + - If the last rule used was `free-varᵀ`, then $$P = x$$, and from + the assumption that $$M$$ is well typed under $$Γ$$ we have + immediately that $$Γ$$ assigns a type to $$x$$. - - If the last rule used was `app1`, then $$t = t_1\;t_2$$ and $$x$$ - appears free in $$t_1$$. Since $$t$$ is well typed under $$\Gamma$$, - we can see from the typing rules that $$t_1$$ must also be, and - the IH then tells us that $$\Gamma$$ assigns $$x$$ a type. + - If the last rule used was `free-·₁`, then $$P = L ·ᵀ M$$ and $$x$$ + appears free in $$L$$. Since $$L$$ is well typed under $$\Gamma$$, + we can see from the typing rules that $$L$$ must also be, and + the IH then tells us that $$Γ$$ assigns $$x$$ a type. - Almost all the other cases are similar: $$x$$ appears free in a - subterm of $$t$$, and since $$t$$ is well typed under $$\Gamma$$, we - know the subterm of $$t$$ in which $$x$$ appears is well typed - under $$\Gamma$$ as well, and the IH gives us exactly the + subterm of $$P$$, and since $$P$$ is well typed under $$Γ$$, we + know the subterm of $$M$$ in which $$x$$ appears is well typed + under $$Γ$$ as well, and the IH gives us exactly the conclusion we want. - - The only remaining case is `abs`. In this case $$t = - \lambda y:A.t'$$, and $$x$$ appears free in $$t'$$; we also know that + - The only remaining case is `free-λᵀ`. In this case $$P = + λᵀ y ∈ A ⇒ N$$, and $$x$$ appears free in $$N$$; we also know that $$x$$ is different from $$y$$. The difference from the previous - cases is that whereas $$t$$ is well typed under $$\Gamma$$, its - body $$t'$$ is well typed under $$(\Gamma, y:A)$$, so the IH + cases is that whereas $$P$$ is well typed under $$\Gamma$$, its + body $$N$$ is well typed under $$(Γ , y ↦ A)$$, so the IH allows us to conclude that $$x$$ is assigned some type by the - extended context $$(\Gamma, y:A)$$. To conclude that $$\Gamma$$ + extended context $$(Γ , y ↦ A)$$. To conclude that $$Γ$$ assigns a type to $$x$$, we appeal the decidable equality for names `_≟_`, noting that $$x$$ and $$y$$ are different variables. -
{% raw %} -freeInCtxt var (var _ x∶A) = (_ , x∶A) -freeInCtxt (app1 x∈t₁) (app t₁∶A _ ) = freeInCtxt x∈t₁ t₁∶A -freeInCtxt (app2 x∈t₂) (app _ t₂∶A) = freeInCtxt x∈t₂ t₂∶A -freeInCtxt (if1 x∈t₁) (if t₁∶A then _ else _ ) = freeInCtxt x∈t₁ t₁∶A -freeInCtxt (if2 x∈t₂) (if _ then t₂∶A else _ ) = freeInCtxt x∈t₂ t₂∶A -freeInCtxt (if3 x∈t₃) (if _ then _ else t₃∶A) = freeInCtxt x∈t₃ t₃∶A -freeInCtxt {x} (abs {y} y≠x x∈t) (abs t∶B) - with freeInCtxt x∈t t∶B -... | x∶A - with y ≟ x -... | yes y=x = ⊥-elim (y≠x y=x) -... | no _ = x∶A -{% endraw %}+
-Next, we'll need the fact that any term $$t$$ which is well typed in +freeLemma free-varᵀ (Ax Γx≡justA) = (_ , Γx≡justA) +freeLemma (free-·ᵀ₁ x∈L) (⇒-E ⊢L ⊢M) = freeLemma x∈L ⊢L +freeLemma (free-·ᵀ₂ x∈M) (⇒-E ⊢L ⊢M) = freeLemma x∈M ⊢M +freeLemma (free-ifᵀ₁ x∈L) (𝔹-E ⊢L ⊢M ⊢N) = freeLemma x∈L ⊢L +freeLemma (free-ifᵀ₂ x∈M) (𝔹-E ⊢L ⊢M ⊢N) = freeLemma x∈M ⊢M +freeLemma (free-ifᵀ₃ x∈N) (𝔹-E ⊢L ⊢M ⊢N) = freeLemma x∈N ⊢N +freeLemma (free-λᵀ {x} {y} y≢x x∈N) (⇒-I ⊢N) with freeLemma x∈N ⊢N +... | Γx=justC with y ≟ x +... | yes y≡x = ⊥-elim (y≢x y≡x) +... | no _ = Γx=justC + ++ +[A subtle point: if the first argument of $$free-λᵀ$$ was of type +$$x ≢ y$$ rather than of type $$y ≢ x$$, then the type of the +term $$Γx=justC$$ would not simplify properly.] + +Next, we'll need the fact that any term $$M$$ which is well typed in the empty context is closed (it has no free variables). #### Exercise: 2 stars, optional (∅⊢-closed) -
{% raw %} - + +postulate ∅⊢-closed : ∀ {t A} → ∅ ⊢ t ∶ A → Closed t -{% endraw %}- -
{% raw %} -∅⊢-closed′ : ∀ {t A} → ∅ ⊢ t ∶ A → Closed t -∅⊢-closed′ (var x ()) -∅⊢-closed′ (app t₁∶A⇒B t₂∶A) (app1 x∈t₁) = ∅⊢-closed′ t₁∶A⇒B x∈t₁ -∅⊢-closed′ (app t₁∶A⇒B t₂∶A) (app2 x∈t₂) = ∅⊢-closed′ t₂∶A x∈t₂ -∅⊢-closed′ true () -∅⊢-closed′ false () -∅⊢-closed′ (if t₁∶bool then t₂∶A else t₃∶A) (if1 x∈t₁) = ∅⊢-closed′ t₁∶bool x∈t₁ -∅⊢-closed′ (if t₁∶bool then t₂∶A else t₃∶A) (if2 x∈t₂) = ∅⊢-closed′ t₂∶A x∈t₂ -∅⊢-closed′ (if t₁∶bool then t₂∶A else t₃∶A) (if3 x∈t₃) = ∅⊢-closed′ t₃∶A x∈t₃ -∅⊢-closed′ (abs {x = x} t′∶A) {y} (abs x≠y y∈t′) with freeInCtxt y∈t′ t′∶A -∅⊢-closed′ (abs {xM A} → =∅ x}⊢ t′∶A)M {y}∈ (abs x≠y y∈t′) | A ,→ y∶Aclosed with x ≟ yM -∅⊢-closed′ (abs {x = x} t′∶A) + +_Proof_: By induction on the derivation of -$$\Gamma \vdash t \in T$$. +$$Γ ⊢ M ∈ A$$. - If the last rule in the derivation was `var`, then $$t = x$$ and $$\Gamma x = T$$. By assumption, $$\Gamma' x = T$$ as well, and @@ -3748,892 +3678,970 @@ $$\Gamma \vdash t \in T$$. $$t_1$$ are also free in $$t_1\;t_2$$, and similarly for $$t_2$$; hence the desired result follows from the induction hypotheses. -+-Sometimes, when we have a proof $$\Gamma\vdash t : A$$, we will need to -replace $$\Gamma$$ by a different context $$\Gamma'$$. When is it safe +Sometimes, when we have a proof $$Γ ⊢ M ∈ A$$, we will need to +replace $$Γ$$ by a different context $$Γ′$$. When is it safe to do this? Intuitively, it must at least be the case that -$$\Gamma'$$ assigns the same types as $$\Gamma$$ to all the variables -that appear free in $$t$$. In fact, this is the only condition that +$$Γ′$$ assigns the same types as $$Γ$$ to all the variables +that appear free in $$M$$. In fact, this is the only condition that is needed. -+ +contradiction {y: }∀ (abs{X : x≠ySet} y∈t′)→ |∀ {x A: ,X} y∶A→ |¬ yes(_≡_ x=y{A = x≠yMaybe x=yX} (just x) nothing) ∅⊢-closed′contradiction (abs {x = x} t′∶A) {y}() + +∅⊢-closed′ (abs x≠y y∈t′): |∀ {M A , () |} → ∅ no _⊢ M ∈ A → closed M -{% endraw %}+∅⊢-closed′ {M} {A} ⊢M {x} x∈M with freeLemma x∈M ⊢M +... | (B , ∅x≡justB) = contradiction (trans (sym ∅x≡justB) (apply-∅ x)) + +{% raw %} -replaceCtxt + +weaken : ∀ {Γ Γ′ t A} - → (∀ {x} → x FreeIn t → Γ x ≡ Γ′ x) - → Γ ⊢ t ∶ A - → Γ′ ⊢∀ t{Γ ∶Γ′ M A} + → (∀ {x} → x FreeIn M → Γ x ≡ Γ′ x) + → Γ ⊢ M ∈ A + → Γ′ ⊢ M ∈ A -{% endraw %}+ +
{% raw %} -replaceCtxt + +weaken fΓ~Γ′ (varAx x x∶AΓx≡justA) rewrite f var = var x x∶A -replaceCtxt f (app t₁∶A⇒B t₂∶A) - = app (replaceCtxt rewrite (fΓ~Γ′ ∘ app1free-varᵀ) t₁∶A⇒B)= (replaceCtxtAx (f ∘ app2) t₂∶A)Γx≡justA replaceCtxtweaken {Γ} {Γ′} {λᵀ x {Γ}∈ A {Γ′⇒ N} fΓ~Γ′ (abs {.Γ}⇒-I ⊢N{x) }= {A}⇒-I {B}(weaken {t′} t′∶B)Γx~Γ′x ⊢N) = abs (replaceCtxtwhere + Γx~Γ′x f′ t′∶B) - where - f′ : ∀ {y} → y FreeIn N → (Γ ∀, {y}x ↦ →A) y FreeIn≡ (Γ′ t′ → (Γ , x ∶↦ A) y ≡ (Γ′ , x ∶ A) y - f′ {y} y∈t′ with x ≟ y - ... | yes _ = refl - ... | no x≠y = f (abs x≠y y∈t′) -replaceCtxt _ true = true -replaceCtxt _ false = false -replaceCtxt f (if t₁∶bool then t₂∶A else t₃∶A) Γx~Γ′x {y} y∈N with x ≟ y + ... | yes refl = if replaceCtxt (f ∘ if1) t₁∶boolrefl - then... replaceCtxt| no x≢y = Γ~Γ′ (ffree-λᵀ ∘x≢y if2y∈N) t₂∶A - elseweaken Γ~Γ′ (⇒-E ⊢L ⊢M) = ⇒-E (weaken (Γ~Γ′ ∘ free-·ᵀ₁) ⊢L) (weaken (Γ~Γ′ ∘ free-·ᵀ₂) ⊢M) +weaken Γ~Γ′ 𝔹-I₁ = 𝔹-I₁ +weaken Γ~Γ′ 𝔹-I₂ = 𝔹-I₂ +weaken Γ~Γ′ (𝔹-E ⊢L ⊢M ⊢N) + = 𝔹-E (weaken (Γ~Γ′ ∘ free-ifᵀ₁) ⊢L) replaceCtxt (fweaken (Γ~Γ′ ∘ if3free-ifᵀ₂) t₃∶A⊢M) (weaken (Γ~Γ′ ∘ free-ifᵀ₃) ⊢N) -{% endraw %}+ +{- +replaceCtxt f (var x x∶A +) rewrite f var = var x x∶A +replaceCtxt f (app t₁∶A⇒B t₂∶A) + = app (replaceCtxt (f ∘ app1) t₁∶A⇒B) (replaceCtxt (f ∘ app2) t₂∶A) +replaceCtxt {Γ} {Γ′} f (abs {.Γ} {x} {A} {B} {t′} t′∶B) + = abs (replaceCtxt f′ t′∶B) + where + f′ : ∀ {y} → y FreeIn t′ → (Γ , x ∶ A) y ≡ (Γ′ , x ∶ A) y + f′ {y} y∈t′ with x ≟ y + ... | yes _ = refl + ... | no x≠y = f (abs x≠y y∈t′) +replaceCtxt _ true = true +replaceCtxt _ false = false +replaceCtxt f (if t₁∶bool then t₂∶A else t₃∶A) + = if replaceCtxt (f ∘ if1) t₁∶bool + then replaceCtxt (f ∘ if2) t₂∶A + else replaceCtxt (f ∘ if3) t₃∶A +-} + + + Now we come to the conceptual heart of the proof that reduction preserves types---namely, the observation that _substitution_ preserves types. Formally, the so-called _Substitution Lemma_ says this: Suppose we -have a term $$t$$ with a free variable $$x$$, and suppose we've been -able to assign a type $$T$$ to $$t$$ under the assumption that $$x$$ has -some type $$U$$. Also, suppose that we have some other term $$v$$ and -that we've shown that $$v$$ has type $$U$$. Then, since $$v$$ satisfies -the assumption we made about $$x$$ when typing $$t$$, we should be -able to substitute $$v$$ for each of the occurrences of $$x$$ in $$t$$ -and obtain a new term that still has type $$T$$. +have a term $$N$$ with a free variable $$x$$, and suppose we've been +able to assign a type $$B$$ to $$N$$ under the assumption that $$x$$ has +some type $$A$$. Also, suppose that we have some other term $$V$$ and +that we've shown that $$V$$ has type $$A$$. Then, since $$V$$ satisfies +the assumption we made about $$x$$ when typing $$N$$, we should be +able to substitute $$V$$ for each of the occurrences of $$x$$ in $$N$$ +and obtain a new term that still has type $$B$$. -_Lemma_: If $$\Gamma,x:U \vdash t : T$$ and $$\vdash v : U$$, then -$$\Gamma \vdash [x:=v]t : T$$. +_Lemma_: If $$Γ , x ↦ A ⊢ N ∈ B$$ and $$∅ ⊢ V ∈ A$$, then +$$Γ ⊢ (N [ x := V ]) ∈ B$$. -
{% raw %} -[:=]-preserves-⊢ + +preservation-[:=] : ∀ {Γ x A tN v B V} → (Γ , x ↦ A) ⊢ N ∈ B + → ∅ ⊢ vV ∶∈ A → Γ , x ∶ A ⊢ t(N ∶ B - → Γ , x ∶ A ⊢ [ x := vV ]) t∈ ∶ B -{% endraw %}+ + One technical subtlety in the statement of the lemma is that -we assign $$v$$ the type $$U$$ in the _empty_ context---in other -words, we assume $$v$$ is closed. This assumption considerably -simplifies the $$abs$$ case of the proof (compared to assuming -$$\Gamma \vdash v : U$$, which would be the other reasonable assumption +we assign $$V$$ the type $$A$$ in the _empty_ context---in other +words, we assume $$V$$ is closed. This assumption considerably +simplifies the $$λᵀ$$ case of the proof (compared to assuming +$$Γ ⊢ V ∈ A$$, which would be the other reasonable assumption at this point) because the context invariance lemma then tells us -that $$v$$ has type $$U$$ in any context at all---we don't have to -worry about free variables in $$v$$ clashing with the variable being -introduced into the context by $$abs$$. +that $$V$$ has type $$A$$ in any context at all---we don't have to +worry about free variables in $$V$$ clashing with the variable being +introduced into the context by $$λᵀ$$. The substitution lemma can be viewed as a kind of "commutation" property. Intuitively, it says that substitution and typing can be done in either order: we can either assign types to the terms -$$t$$ and $$v$$ separately (under suitable contexts) and then combine +$$N$$ and $$V$$ separately (under suitable contexts) and then combine them using substitution, or we can substitute first and then -assign a type to $$ $$x:=v$$ t $$---the result is the same either +assign a type to $$N [ x := V ]$$---the result is the same either way. -_Proof_: We show, by induction on $$t$$, that for all $$T$$ and -$$\Gamma$$, if $$\Gamma,x:U \vdash t : T$$ and $$\vdash v : U$$, then $$\Gamma -\vdash $$x:=v$$t : T$$. +_Proof_: We show, by induction on $$N$$, that for all $$A$$ and +$$Γ$$, if $$Γ , x ↦ A \vdash N ∈ B$$ and $$∅ ⊢ V ∈ A$$, then +$$Γ \vdash N [ x := V ] ∈ B$$. - - If $$t$$ is a variable there are two cases to consider, - depending on whether $$t$$ is $$x$$ or some other variable. + - If $$N$$ is a variable there are two cases to consider, + depending on whether $$N$$ is $$x$$ or some other variable. - - If $$t = x$$, then from the fact that $$\Gamma, x:U \vdash x : - T$$ we conclude that $$U = T$$. We must show that $$[x:=v]x = - v$$ has type $$T$$ under $$\Gamma$$, given the assumption that - $$v$$ has type $$U = T$$ under the empty context. This + - If $$N = varᵀ x$$, then from the fact that $$Γ , x ↦ A ⊢ N ∈ B$$ + we conclude that $$A = B$$. We must show that $$x [ x := V] = + V$$ has type $$A$$ under $$Γ$$, given the assumption that + $$V$$ has type $$A$$ under the empty context. This follows from context invariance: if a closed term has type - $$T$$ in the empty context, it has that type in any context. + $$A$$ in the empty context, it has that type in any context. - - If $$t$$ is some variable $$y$$ that is not equal to $$x$$, then - we need only note that $$y$$ has the same type under $$\Gamma, - x:U$$ as under $$\Gamma$$. + - If $$N$$ is some variable $$x′$$ different from $$x$$, then + we need only note that $$x′$$ has the same type under $$Γ , x ↦ A$$ + as under $$Γ$$. - - If $$t$$ is an abstraction $$\lambda y:t_{11}. t_{12}$$, then the IH tells us, - for all $$\Gamma'$$ and $$T'$$, that if $$\Gamma',x:U \vdash t_{12}:T'$$ - and $$\vdash v:U$$, then $$\Gamma' \vdash [x:=v]t_{12}:T'$$. + - If $$N$$ is an abstraction $$λᵀ x′ ∈ A′ ⇒ N′$$, then the IH tells us, + for all $$Γ′$$́ and $$B′$$, that if $$Γ′ , x ↦ A ⊢ N′ ∈ B′$$ + and $$∅ ⊢ V ∈ A$$, then $$Γ′ ⊢ N′ [ x := V ] ∈ B′$$. The substitution in the conclusion behaves differently - depending on whether $$x$$ and $$y$$ are the same variable. + depending on whether $$x$$ and $$x′$$ are the same variable. - First, suppose $$x = y$$. Then, by the definition of - substitution, $$[x:=v]t = t$$, so we just need to show $$\Gamma \vdash - t : T$$. But we know $$\Gamma,x:U \vdash t : T$$, and, since $$y$$ - does not appear free in $$\lambda y:t_{11}. t_{12}$$, the context invariance - lemma yields $$\Gamma \vdash t : T$$. + First, suppose $$x ≡ x′$$. Then, by the definition of + substitution, $$N [ x := V] = N$$, so we just need to show $$Γ ⊢ N ∈ B$$. + But we know $$Γ , x ↦ A ⊢ N ∈ B$$ and, since $$x ≡ x′$$ + does not appear free in $$λᵀ x′ ∈ A′ ⇒ N′$$, the context invariance + lemma yields $$Γ ⊢ N ∈ B$$. - Second, suppose $$x \neq y$$. We know $$\Gamma,x:U,y:t_{11} \vdash - t_{12}:t_{12}$$ by inversion of the typing relation, from which - $$\Gamma,y:t_{11},x:U \vdash t_{12}:t_{12}$$ follows by the context invariance - lemma, so the IH applies, giving us $$\Gamma,y:t_{11} \vdash - [x:=v]t_{12}:t_{12}$$. By $$abs$$, $$\Gamma \vdash \lambda y:t_{11}. - [x:=v]t_{12}:t_{11}\to t_{12}$$, and by the definition of substitution (noting - that $$x \neq y$$), $$\Gamma \vdash \lambda y:t_{11}. [x:=v]t_{12}:t_{11}\to - t_{12}$$ as required. + Second, suppose $$x ≢ x′$$. We know $$Γ , x ↦ A , x′ ↦ A′ ⊢ N′ ∈ B′$$ + by inversion of the typing relation, from which + $$Γ , x′ ↦ A′ , x ↦ A ⊢ N′ ∈ B′$$ follows by update permute, + so the IH applies, giving us $$Γ , x′ ↦ A′ ⊢ N′ [ x := V ] ∈ B′$$ + By $$⇒-I$$, we have $$Γ ⊢ λᵀ x′ ∈ A′ ⇒ (N′ [ x := V ]) ∈ A′ ⇒ B′$$ + and the definition of substitution (noting $$x ≢ x′$$) gives + $$Γ ⊢ (λᵀ x′ ∈ A′ ⇒ N′) [ x := V ] ∈ A′ ⇒ B′$$ as required. - - If $$t$$ is an application $$t_1 t_2$$, the result follows + - If $$N$$ is an application $$L′ ·ᵀ M′$$, the result follows straightforwardly from the definition of substitution and the induction hypotheses. - The remaining cases are similar to the application case. -One more technical note: This proof is a rare case where an -induction on terms, rather than typing derivations, yields a -simpler argument. The reason for this is that the assumption -$$update Gamma x U \vdash t : T$$ is not completely generic, in the -sense that one of the "slots" in the typing relation---namely the -context---is not just a variable, and this means that Agda's -native induction tactic does not give us the induction hypothesis -that we want. It is possible to work around this, but the needed -generalization is a little tricky. The term $$t$$, on the other -hand, _is_ completely generic. +We need a couple of lemmas. A closed term can be weakened to any context, and just is injective. +
-{% raw %} -[:=]-preserves-⊢weaken-closed : ∀ {V A Γ} → ∅ ⊢ V ∈ A → Γ ⊢ V ∈ A +weaken-closed {V} {A} } {x}⊢V v∶A= (varweaken yΓ~Γ′ y∈Γ)⊢V with x ≟ y -...where + Γ~Γ′ |: yes∀ x=y{x} → =x {!!} -...FreeIn V |→ no x≠y∅ x ≡ Γ =x {!!} -[:=]-preserves-⊢Γ~Γ′ v∶A{x} (absx∈V t′∶B) = {!!}⊥-elim (x∉V x∈V) + where -[:=]-preserves-⊢x∉V v∶A: ¬ (appx t₁∶A⇒BFreeIn V) + x∉V t₂∶A) = ∅⊢-closed ⊢V {x} - appjust-injective ([:=]-preserves-⊢: v∶A∀ {X t₁∶A⇒B): ([:=]-preserves-⊢Set} v∶A{x t₂∶A) -[:=]-preserves-⊢y v∶A: true X} → _≡_ {A = Maybe X} (just x) (just y) → x true≡ y [:=]-preserves-⊢just-injective v∶Arefl = false = falserefl -[:=]-preserves-⊢ + +### Main Theorem We now have the tools we need to prove preservation: if a closed -term $$t$$ has type $$T$$ and takes a step to $$t'$$, then $$t'$$ -is also a closed term with type $$T$$. In other words, the small-step -reduction relation preserves types. +term $$M$$ has type $$A$$ and takes a step to $$N$$, then $$N$$ +is also a closed term with type $$A$$. In other words, small-step +reduction preserves types. -Theorem preservation : forall t t' T, - empty \vdash t : T → - t ==> t' → - empty \vdash t' : T. ++ +preservation-[:=] v∶A (if t₁∶bool then{_} t₂∶B{x} else(Ax t₃∶B{_} {x′} [Γ,x↦A]x′≡B) = - if [:=]-preserves-⊢⊢V with x ≟ v∶A t₁∶boolx′ - then...| [:=]-preserves-⊢yes v∶Ax≡x′ rewrite just-injective t₂∶B[Γ,x↦A]x′≡B = weaken-closed ⊢V +...| no x≢x′ = Ax [Γ,x↦A]x′≡B +preservation-[:=] {Γ} {x} {A} {λᵀ x′ ∈ A′ ⇒ N′} {.A′ ⇒ B′} {V} (⇒-I ⊢N′) ⊢V with x ≟ x′ +...| yes x≡x′ rewrite x≡x′ = weaken Γ′~Γ (⇒-I ⊢N′) else [:=]-preserves-⊢ v∶A t₃∶Bwhere -{% endraw %}+ Γ′~Γ : ∀ {y} → y FreeIn (λᵀ x′ ∈ A′ ⇒ N′) → (Γ , x′ ↦ A) y ≡ Γ y + Γ′~Γ {y} (free-λᵀ x′≢y y∈N′) with x′ ≟ y + ...| yes x′≡y = ⊥-elim (x′≢y x′≡y) + ...| no _ = refl +...| no x≢x′ = ⇒-I ⊢N′V + where + x′x⊢N′ : (Γ , x′ ↦ A′ , x ↦ A) ⊢ N′ ∈ B′ + x′x⊢N′ rewrite update-permute Γ x A x′ A′ x≢x′ = ⊢N′ + ⊢N′V : (Γ , x′ ↦ A′) ⊢ N′ [ x := V ] ∈ B′ + ⊢N′V = preservation-[:=] x′x⊢N′ ⊢V +preservation-[:=] (⇒-E ⊢L ⊢M) ⊢V = ⇒-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) +preservation-[:=] 𝔹-I₁ ⊢V = 𝔹-I₁ +preservation-[:=] 𝔹-I₂ ⊢V = 𝔹-I₂ +preservation-[:=] (𝔹-E ⊢L ⊢M ⊢N) ⊢V = + 𝔹-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) (preservation-[:=] ⊢N ⊢V) + ++ +preservation : ∀ {M N A} → ∅ ⊢ M ∈ A → M ⟹ N → ∅ ⊢ N ∈ A + +_Proof_: By induction on the derivation of $$\vdash t : T$$. @@ -5012,6 +6024,458 @@ _Proof_: By induction on the derivation of $$\vdash t : T$$. - Otherwise, $$t$$ steps by $$Sif$$, and the desired conclusion follows directly from the induction hypothesis. ++ +preservation (Ax x₁) () +preservation (⇒-I ⊢N) () +preservation (⇒-E (⇒-I ⊢N) ⊢V) (β⇒ valueV) = preservation-[:=] ⊢N ⊢V +preservation (⇒-E ⊢L ⊢M) (γ⇒₁ L⟹L′) with preservation ⊢L L⟹L′ +...| ⊢L′ = ⇒-E ⊢L′ ⊢M +preservation (⇒-E ⊢L ⊢M) (γ⇒₂ valueL M⟹M′) with preservation ⊢M M⟹M′ +...| ⊢M′ = ⇒-E ⊢L ⊢M′ +preservation 𝔹-I₁ () +preservation 𝔹-I₂ () +preservation (𝔹-E 𝔹-I₁ ⊢M ⊢N) β𝔹₁ = ⊢M +preservation (𝔹-E 𝔹-I₂ ⊢M ⊢N) β𝔹₂ = ⊢N +preservation (𝔹-E ⊢L ⊢M ⊢N) (γ𝔹 L⟹L′) with preservation ⊢L L⟹L′ +...| ⊢L′ = 𝔹-E ⊢L′ ⊢M ⊢N + +-- Writing out implicit parameters in full +-- preservation (⇒-E {Γ} {λᵀ x ∈ A ⇒ N} {M} {.A} {B} (⇒-I {.Γ} {.x} {.N} {.A} {.B} ⊢N) ⊢M) (β⇒ {.x} {.A} {.N} {.M} valueM) +-- = preservation-[:=] {Γ} {x} {A} {M} {N} {B} ⊢M ⊢N + ++ Proof with eauto. remember (@empty ty) as Gamma. intros t t' T HT. generalize dependent t'. @@ -5034,7 +6498,6 @@ expressions. Does this property hold for STLC? That is, is it always the case that, if $$t ==> t'$$ and $$has_type t' T$$, then $$empty \vdash t : T$$? If so, prove it. If not, give a counter-example not involving conditionals. -(* FILL IN HERE ## Type Soundness @@ -5053,8 +6516,6 @@ Proof. intros t t' T Hhas_type Hmulti. unfold stuck. intros $$Hnf Hnot_val$$. unfold normal_form in Hnf. induction Hmulti. - (* FILL IN HERE Admitted. -(** $$$$ ## Uniqueness of Types @@ -5064,9 +6525,6 @@ Another nice property of the STLC is that types are unique: a given term (in a given context) has at most one type. Formalize this statement and prove it. -(* FILL IN HERE -(** $$$$ - ## Additional Exercises @@ -5256,3 +6714,4 @@ with arithmetic. Specifically: - Extend the proofs of all the properties (up to $$soundness$$) of the original STLC to deal with the new syntactic forms. Make sure Agda accepts the whole file. + diff --git a/src/StlcProp.lagda b/src/StlcProp.lagda index 0980a17d..6f9a40d8 100644 --- a/src/StlcProp.lagda +++ b/src/StlcProp.lagda @@ -483,12 +483,7 @@ just-injective refl = refl preservation-[:=] {_} {x} (Ax {_} {x′} [Γ,x↦A]x′≡B) ⊢V with x ≟ x′ ...| yes x≡x′ rewrite just-injective [Γ,x↦A]x′≡B = weaken-closed ⊢V ...| no x≢x′ = Ax [Γ,x↦A]x′≡B -{- -preservation-[:=] {Γ} {x} {A} {varᵀ x′} {B} {V} (Ax {.(Γ , x ↦ A)} {.x′} {.B} Γx′≡B) ⊢V with x ≟ x′ -...| yes x≡x′ rewrite just-injective Γx′≡B = weaken-closed ⊢V -...| no x≢x′ = Ax {Γ} {x′} {B} Γx′≡B --} -preservation-[:=] {Γ} {x} {A} {λᵀ x′ ∈ A′ ⇒ N′} {.A′ ⇒ B′} {V} (⇒-I {.(Γ , x ↦ A)} {.x′} {.N′} {.A′} {.B′} ⊢N′) ⊢V with x ≟ x′ +preservation-[:=] {Γ} {x} {A} {λᵀ x′ ∈ A′ ⇒ N′} {.A′ ⇒ B′} {V} (⇒-I ⊢N′) ⊢V with x ≟ x′ ...| yes x≡x′ rewrite x≡x′ = weaken Γ′~Γ (⇒-I ⊢N′) where Γ′~Γ : ∀ {y} → y FreeIn (λᵀ x′ ∈ A′ ⇒ N′) → (Γ , x′ ↦ A) y ≡ Γ y @@ -498,35 +493,14 @@ preservation-[:=] {Γ} {x} {A} {λᵀ x′ ∈ A′ ⇒ N′} {.A′ ⇒ B′} { ...| no x≢x′ = ⇒-I ⊢N′V where x′x⊢N′ : (Γ , x′ ↦ A′ , x ↦ A) ⊢ N′ ∈ B′ - x′x⊢N′ rewrite update-permute Γ x A x′ A′ x≢x′ = {!⊢N′!} + x′x⊢N′ rewrite update-permute Γ x A x′ A′ x≢x′ = ⊢N′ ⊢N′V : (Γ , x′ ↦ A′) ⊢ N′ [ x := V ] ∈ B′ ⊢N′V = preservation-[:=] x′x⊢N′ ⊢V -{- -...| yes x′≡x rewrite x′≡x | update-shadow Γ x A A′ = {!!} - -- ⇒-I ⊢N′ -...| no x′≢x rewrite update-permute Γ x′ A′ x A x′≢x = {!!} - -- ⇒-I {Γ} {x′} {N′} {A′} {B′} (preservation-[:=] {(Γ , x′ ↦ A′)} {x} {A} ⊢N′ ⊢V) --} preservation-[:=] (⇒-E ⊢L ⊢M) ⊢V = ⇒-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) preservation-[:=] 𝔹-I₁ ⊢V = 𝔹-I₁ preservation-[:=] 𝔹-I₂ ⊢V = 𝔹-I₂ preservation-[:=] (𝔹-E ⊢L ⊢M ⊢N) ⊢V = 𝔹-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) (preservation-[:=] ⊢N ⊢V) - -{- -[:=]-preserves-⊢ {Γ} {x} v∶A (var y y∈Γ) with x ≟ y -... | yes x=y = {!!} -... | no x≠y = {!!} -[:=]-preserves-⊢ v∶A (abs t′∶B) = {!!} -[:=]-preserves-⊢ v∶A (app t₁∶A⇒B t₂∶A) = - app ([:=]-preserves-⊢ v∶A t₁∶A⇒B) ([:=]-preserves-⊢ v∶A t₂∶A) -[:=]-preserves-⊢ v∶A true = true -[:=]-preserves-⊢ v∶A false = false -[:=]-preserves-⊢ v∶A (if t₁∶bool then t₂∶B else t₃∶B) = - if [:=]-preserves-⊢ v∶A t₁∶bool - then [:=]-preserves-⊢ v∶A t₂∶B - else [:=]-preserves-⊢ v∶A t₃∶B --} \end{code}