--- title : "Properties: Progress and Preservation" permalink : /Properties/ --- ```agda module plfa.part2.Properties where ``` This chapter covers properties of the simply-typed lambda calculus, as introduced in the previous chapter. The most important of these properties are progress and preservation. We introduce these below, and show how to combine them to get Agda to compute reduction sequences for us. ## Imports ```agda open import Relation.Binary.PropositionalEquality using (_≡_; _≢_; refl; sym; cong; cong₂) open import Data.String using (String; _≟_) open import Data.Nat using (ℕ; zero; suc) open import Data.Empty using (⊥; ⊥-elim) open import Data.Product using (_×_; proj₁; proj₂; ∃; ∃-syntax) renaming (_,_ to ⟨_,_⟩) open import Data.Sum using (_⊎_; inj₁; inj₂) open import Relation.Nullary using (¬_; Dec; yes; no) open import Function using (_∘_) open import plfa.part1.Isomorphism open import plfa.part2.Lambda ``` ## Introduction The last chapter introduced simply-typed lambda calculus, including the notions of closed terms, terms that are values, reducing one term to another, and well-typed terms. Ultimately, we would like to show that we can keep reducing a term until we reach a value. For instance, in the last chapter we showed that two plus two is four, plus · two · two —↠ `suc `suc `suc `suc `zero which was proved by a long chain of reductions, ending in the value on the right. Every term in the chain had the same type, `` `ℕ ``. We also saw a second, similar example involving Church numerals. What we might expect is that every term is either a value or can take a reduction step. As we will see, this property does _not_ hold for every term, but it does hold for every closed, well-typed term. _Progress_: If `∅ ⊢ M ⦂ A` then either `M` is a value or there is an `N` such that `M —→ N`. So, either we have a value, and we are done, or we can take a reduction step. In the latter case, we would like to apply progress again. But to do so we need to know that the term yielded by the reduction is itself closed and well typed. It turns out that this property holds whenever we start with a closed, well-typed term. _Preservation_: If `∅ ⊢ M ⦂ A` and `M —→ N` then `∅ ⊢ N ⦂ A`. This gives us a recipe for automating evaluation. Start with a closed and well-typed term. By progress, it is either a value, in which case we are done, or it reduces to some other term. By preservation, that other term will itself be closed and well typed. Repeat. We will either loop forever, in which case evaluation does not terminate, or we will eventually reach a value, which is guaranteed to be closed and of the same type as the original term. We will turn this recipe into Agda code that can compute for us the reduction sequence of `plus · two · two`, and its Church numeral variant. (The development in this chapter was inspired by the corresponding development in _Software Foundations_, Volume _Programming Language Foundations_, Chapter _StlcProp_. It will turn out that one of our technical choices — to introduce an explicit judgment `Γ ∋ x ⦂ A` in place of treating a context as a function from identifiers to types — permits a simpler development. In particular, we can prove substitution preserves types without needing to develop a separate inductive definition of the `appears_free_in` relation.) ## Values do not reduce We start with an easy observation. Values do not reduce: ```agda V¬—→ : ∀ {M N} → Value M ---------- → ¬ (M —→ N) V¬—→ V-ƛ () V¬—→ V-zero () V¬—→ (V-suc VM) (ξ-suc M—→N) = V¬—→ VM M—→N ``` We consider the three possibilities for values: * If it is an abstraction then no reduction applies * If it is zero then no reduction applies * If it is a successor then rule `ξ-suc` may apply, but in that case the successor is itself of a value that reduces, which by induction cannot occur. As a corollary, terms that reduce are not values: ```agda —→¬V : ∀ {M N} → M —→ N --------- → ¬ Value M —→¬V M—→N VM = V¬—→ VM M—→N ``` If we expand out the negations, we have V¬—→ : ∀ {M N} → Value M → M —→ N → ⊥ —→¬V : ∀ {M N} → M —→ N → Value M → ⊥ which are the same function with the arguments swapped. #### Exercise `Canonical-≃` (practice) Well-typed values must take one of a small number of _canonical forms_, which provide an analogue of the `Value` relation that relates values to their types. A lambda expression must have a function type, and a zero or successor expression must be a natural. Further, the body of a function must be well typed in a context containing only its bound variable, and the argument of successor must itself be canonical: ```agda infix 4 Canonical_⦂_ data Canonical_⦂_ : Term → Type → Set where C-ƛ : ∀ {x A N B} → ∅ , x ⦂ A ⊢ N ⦂ B ----------------------------- → Canonical (ƛ x ⇒ N) ⦂ (A ⇒ B) C-zero : -------------------- Canonical `zero ⦂ `ℕ C-suc : ∀ {V} → Canonical V ⦂ `ℕ --------------------- → Canonical `suc V ⦂ `ℕ ``` Show that `Canonical V ⦂ A` is isomorphic to `(∅ ⊢ V ⦂ A) × (Value V)`, that is, the canonical forms are exactly the well-typed values. ``` open import plfa.part1.Isomorphism using (_≃_) module Can-Iso where to : ∀ {V A} → Canonical V ⦂ A → ∅ ⊢ V ⦂ A × Value V to (C-ƛ ⊢N) = ⟨ ⊢ƛ ⊢N , V-ƛ ⟩ to C-zero = ⟨ ⊢zero , V-zero ⟩ to (C-suc can-V) = let ⟨ ⊢V , val-V ⟩ = to can-V in ⟨ (⊢suc ⊢V) , (V-suc val-V) ⟩ from : ∀ {V A} → ∅ ⊢ V ⦂ A × Value V → Canonical V ⦂ A from ⟨ ⊢ƛ ⊢V , V-ƛ ⟩ = C-ƛ ⊢V from ⟨ ⊢zero , V-zero ⟩ = C-zero from ⟨ ⊢suc ⊢V , V-suc val-V ⟩ = C-suc (from ⟨ ⊢V , val-V ⟩) from∘to : ∀ {V A} → (x : Canonical V ⦂ A) → from (to x) ≡ x from∘to (C-ƛ ⊢N) = refl from∘to C-zero = refl from∘to (C-suc can-V) = cong C-suc (from∘to can-V) to∘from : ∀ {V A} → (y : ∅ ⊢ V ⦂ A × Value V) → to (from y) ≡ y to∘from ⟨ ⊢ƛ ⊢V , V-ƛ ⟩ = refl to∘from ⟨ ⊢zero , V-zero ⟩ = refl to∘from ⟨ ⊢suc ⊢V , V-suc val-V ⟩ rewrite to∘from ⟨ ⊢V , val-V ⟩ = refl Canonical-≃ : ∀ {V A} → Canonical V ⦂ A ≃ (∅ ⊢ V ⦂ A) × (Value V) Canonical-≃ = record { to = Can-Iso.to ; from = Can-Iso.from ; from∘to = Can-Iso.from∘to ; to∘from = Can-Iso.to∘from } ``` ## Progress We would like to show that every term is either a value or takes a reduction step. However, this is not true in general. The term `zero · `suc `zero is neither a value nor can take a reduction step. And if `` "s" ⦂ `ℕ ⇒ `ℕ `` then the term ` "s" · `zero cannot reduce because we do not know which function is bound to the free variable `"s"`. The first of these terms is ill typed, and the second has a free variable. Every term that is well typed and closed has the desired property. _Progress_: If `∅ ⊢ M ⦂ A` then either `M` is a value or there is an `N` such that `M —→ N`. To formulate this property, we first introduce a relation that captures what it means for a term `M` to make progress: ```agda data Progress (M : Term) : Set where step : ∀ {N} → M —→ N ---------- → Progress M done : Value M ---------- → Progress M ``` A term `M` makes progress if either it can take a step, meaning there exists a term `N` such that `M —→ N`, or if it is done, meaning that `M` is a value. If a term is well typed in the empty context then it satisfies progress: ```agda progress : ∀ {M A} → ∅ ⊢ M ⦂ A ---------- → Progress M progress (⊢` ()) progress (⊢ƛ ⊢N) = done V-ƛ progress (⊢L · ⊢M) with progress ⊢L ... | step L—→L′ = step (ξ-·₁ L—→L′) ... | done V-ƛ with progress ⊢M ... | step M—→M′ = step (ξ-·₂ V-ƛ M—→M′) ... | done VM = step (β-ƛ VM) progress ⊢zero = done V-zero progress (⊢suc ⊢M) with progress ⊢M ... | step M—→M′ = step (ξ-suc M—→M′) ... | done VM = done (V-suc VM) progress (⊢case ⊢L ⊢M ⊢N) with progress ⊢L ... | step L—→L′ = step (ξ-case L—→L′) ... | done (V-zero) = step β-zero ... | done (V-suc VL) = step (β-suc VL) progress (⊢μ ⊢M) = step β-μ ``` We induct on the evidence that the term is well typed. Let's unpack the first three cases: * The term cannot be a variable, since no variable is well typed in the empty context. * If the term is a lambda abstraction then it is a value. * If the term is an application `L · M`, recursively apply progress to the derivation that `L` is well typed: + If the term steps, we have evidence that `L —→ L′`, which by `ξ-·₁` means that our original term steps to `L′ · M` + If the term is done, we have evidence that `L` is a value, which must be a lambda abstraction. Recursively apply progress to the derivation that `M` is well typed: - If the term steps, we have evidence that `M —→ M′`, which by `ξ-·₂` means that our original term steps to `L · M′`. Step `ξ-·₂` applies only if we have evidence that `L` is a value, but progress on that subterm has already supplied the required evidence. - If the term is done, we have evidence that `M` is a value, so our original term steps by `β-ƛ`. The remaining cases are similar. If by induction we have a `step` case we apply a `ξ` rule, and if we have a `done` case then either we have a value or apply a `β` rule. For fixpoint, no induction is required as the `β` rule applies immediately. Our code reads neatly in part because we consider the `step` option before the `done` option. We could, of course, do it the other way around, but then the `...` abbreviation no longer works, and we will need to write out all the arguments in full. In general, the rule of thumb is to consider the easy case (here `step`) before the hard case (here `done`). If you have two hard cases, you will have to expand out `...` or introduce subsidiary functions. Instead of defining a data type for `Progress M`, we could have formulated progress using disjunction and existentials: ```agda -- postulate -- progress′ : ∀ M {A} → ∅ ⊢ M ⦂ A → Value M ⊎ ∃[ N ](M —→ N) ``` This leads to a less perspicuous proof. Instead of the mnemonic `done` and `step` we use `inj₁` and `inj₂`, and the term `N` is no longer implicit and so must be written out in full. In the case for `β-ƛ` this requires that we match against the lambda expression `L` to determine its bound variable and body, `ƛ x ⇒ N`, so we can show that `L · M` reduces to `N [ x := M ]`. #### Exercise `Progress-≃` (practice) Show that `Progress M` is isomorphic to `Value M ⊎ ∃[ N ](M —→ N)`. ```agda Progress-≃ : ∀ {M} → Progress M ≃ Value M ⊎ ∃[ N ](M —→ N) Progress-≃ = record { to = λ where (done val-M) → inj₁ val-M (step {N} M—→N) → inj₂ ⟨ N , M—→N ⟩ ; from = λ where (inj₁ val-M) → done val-M (inj₂ ⟨ N , M—→N ⟩) → step M—→N ; from∘to = λ where (done val-M) → refl (step {N} M—→N) → refl ; to∘from = λ where (inj₁ val-M) → refl (inj₂ ⟨ N , M—→N ⟩) → refl } ``` #### Exercise `progress′` (practice) Write out the proof of `progress′` in full, and compare it to the proof of `progress` above. ```agda -- We get progress′ by replacing in progress -- - each `done val-M` with `inj₁ val-M`; and -- - each `step M—→N` with `inj₂ ⟨ _ , M—→N ⟩`. -- This makes it easier to read the type signature, but harder to read -- the constructors. -- Structurally, the proof remains the same, because the `Progress M` -- data-type and `Value M ⊎ ∃[ N ](M —→ N)` are trivially isomorphic (the -- translation between them renames the constructors). progress′ : ∀ {M A} → ∅ ⊢ M ⦂ A ------------------------ → Value M ⊎ ∃[ N ](M —→ N) progress′ (⊢` ()) progress′ (⊢ƛ ⊢N) = inj₁ V-ƛ progress′ (⊢L · ⊢M) with progress′ ⊢L ... | inj₂ ⟨ _ , L—→L′ ⟩ = inj₂ ⟨ _ , (ξ-·₁ L—→L′) ⟩ ... | inj₁ V-ƛ with progress′ ⊢M ... | inj₂ ⟨ _ , M—→M′ ⟩ = inj₂ ⟨ _ , (ξ-·₂ V-ƛ M—→M′) ⟩ ... | inj₁ VM = inj₂ ⟨ _ , (β-ƛ VM) ⟩ progress′ ⊢zero = inj₁ V-zero progress′ (⊢suc ⊢M) with progress′ ⊢M ... | inj₂ ⟨ _ , M—→M′ ⟩ = inj₂ ⟨ _ , (ξ-suc M—→M′) ⟩ ... | inj₁ VM = inj₁ (V-suc VM) progress′ (⊢case ⊢L ⊢M ⊢N) with progress′ ⊢L ... | inj₂ ⟨ _ , L—→L′ ⟩ = inj₂ ⟨ _ , (ξ-case L—→L′) ⟩ ... | inj₁ (V-zero) = inj₂ ⟨ _ , β-zero ⟩ ... | inj₁ (V-suc VL) = inj₂ ⟨ _ , (β-suc VL) ⟩ progress′ (⊢μ ⊢M) = inj₂ ⟨ _ , β-μ ⟩ ``` #### Exercise `value?` (practice) Combine `progress` and `—→¬V` to write a program that decides whether a well-typed term is a value: ```agda value? : ∀ {A M} → ∅ ⊢ M ⦂ A → Dec (Value M) value? (⊢ƛ ⊢M) = yes V-ƛ value? (⊢M₁ · ⊢M₂) = no (λ ()) value? ⊢zero = yes V-zero value? (⊢case ⊢M ⊢M₁ ⊢M₂) = no (λ ()) value? (⊢μ ⊢M) = no (λ ()) value? (⊢suc ⊢M) with value? ⊢M ... | yes val-M = yes (V-suc val-M) ... | no ¬val-M = no (λ { (V-suc val-M) → ¬val-M val-M }) ``` ## Prelude to preservation The other property we wish to prove, preservation of typing under reduction, turns out to require considerably more work. The proof has three key steps. The first step is to show that types are preserved by _renaming_. _Renaming_: Let `Γ` and `Δ` be two contexts such that every variable that appears in `Γ` also appears with the same type in `Δ`. Then if any term is typeable under `Γ`, it has the same type under `Δ`. In symbols: ∀ {x A} → Γ ∋ x ⦂ A → Δ ∋ x ⦂ A --------------------------------- ∀ {M A} → Γ ⊢ M ⦂ A → Δ ⊢ M ⦂ A Three important corollaries follow. The _weaken_ lemma asserts that a term which is well typed in the empty context is also well typed in an arbitrary context. The _drop_ lemma asserts that a term which is well typed in a context where the same variable appears twice remains well typed if we drop the shadowed occurrence. The _swap_ lemma asserts that a term which is well typed in a context remains well typed if we swap two variables. (Renaming is similar to the _context invariance_ lemma in _Software Foundations_, but it does not require the definition of `appears_free_in` nor the `free_in_context` lemma.) The second step is to show that types are preserved by _substitution_. _Substitution_: Say we have a closed term `V` of type `A`, and under the assumption that `x` has type `A` the term `N` has type `B`. Then substituting `V` for `x` in `N` yields a term that also has type `B`. In symbols: ∅ ⊢ V ⦂ A Γ , x ⦂ A ⊢ N ⦂ B -------------------- Γ ⊢ N [ x := V ] ⦂ B The result does not depend on `V` being a value, but it does require that `V` be closed; recall that we restricted our attention to substitution by closed terms in order to avoid the need to rename bound variables. The term into which we are substituting is typed in an arbitrary context `Γ`, extended by the variable `x` for which we are substituting; and the result term is typed in `Γ`. The lemma establishes that substitution composes well with typing: typing the components separately guarantees that the result of combining them is also well typed. The third step is to show preservation. _Preservation_: If `∅ ⊢ M ⦂ A` and `M —→ N` then `∅ ⊢ N ⦂ A`. The proof is by induction over the possible reductions, and the substitution lemma is crucial in showing that each of the `β` rules that uses substitution preserves types. We now proceed with our three-step programme. ## Renaming We often need to "rebase" a type derivation, replacing a derivation `Γ ⊢ M ⦂ A` by a related derivation `Δ ⊢ M ⦂ A`. We may do so as long as every variable that appears in `Γ` also appears in `Δ`, and with the same type. Three of the rules for typing (lambda abstraction, case on naturals, and fixpoint) have hypotheses that extend the context to include a bound variable. In each of these rules, `Γ` appears in the conclusion and `Γ , x ⦂ A` appears in a hypothesis. Thus: Γ , x ⦂ A ⊢ N ⦂ B ------------------- ⊢ƛ Γ ⊢ ƛ x ⇒ N ⦂ A ⇒ B for lambda expressions, and similarly for case and fixpoint. To deal with this situation, we first prove a lemma showing that if one context maps to another, this is still true after adding the same variable to both contexts: ```agda ext : ∀ {Γ Δ} → (∀ {x A} → Γ ∋ x ⦂ A → Δ ∋ x ⦂ A) ----------------------------------------------------- → (∀ {x y A B} → Γ , y ⦂ B ∋ x ⦂ A → Δ , y ⦂ B ∋ x ⦂ A) ext ρ Z = Z ext ρ (S x≢y ∋x) = S x≢y (ρ ∋x) ``` Let `ρ` be the name of the map that takes evidence that `x` appears in `Γ` to evidence that `x` appears in `Δ`. The proof is by case analysis of the evidence that `x` appears in the extended context `Γ , y ⦂ B`: * If `x` is the same as `y`, we used `Z` to access the last variable in the extended `Γ`; and can similarly use `Z` to access the last variable in the extended `Δ`. * If `x` differs from `y`, then we used `S` to skip over the last variable in the extended `Γ`, where `x≢y` is evidence that `x` and `y` differ, and `∋x` is the evidence that `x` appears in `Γ`; and we can similarly use `S` to skip over the last variable in the extended `Δ`, applying `ρ` to find the evidence that `x` appears in `Δ`. With the extension lemma under our belts, it is straightforward to prove renaming preserves types: ```agda rename : ∀ {Γ Δ} → (∀ {x A} → Γ ∋ x ⦂ A → Δ ∋ x ⦂ A) ---------------------------------- → (∀ {M A} → Γ ⊢ M ⦂ A → Δ ⊢ M ⦂ A) rename ρ (⊢` ∋w) = ⊢` (ρ ∋w) rename ρ (⊢ƛ ⊢N) = ⊢ƛ (rename (ext ρ) ⊢N) rename ρ (⊢L · ⊢M) = (rename ρ ⊢L) · (rename ρ ⊢M) rename ρ ⊢zero = ⊢zero rename ρ (⊢suc ⊢M) = ⊢suc (rename ρ ⊢M) rename ρ (⊢case ⊢L ⊢M ⊢N) = ⊢case (rename ρ ⊢L) (rename ρ ⊢M) (rename (ext ρ) ⊢N) rename ρ (⊢μ ⊢M) = ⊢μ (rename (ext ρ) ⊢M) ``` As before, let `ρ` be the name of the map that takes evidence that `x` appears in `Γ` to evidence that `x` appears in `Δ`. We induct on the evidence that `M` is well typed in `Γ`. Let's unpack the first three cases: * If the term is a variable, then applying `ρ` to the evidence that the variable appears in `Γ` yields the corresponding evidence that the variable appears in `Δ`. * If the term is a lambda abstraction, use the previous lemma to extend the map `ρ` suitably and use induction to rename the body of the abstraction. * If the term is an application, use induction to rename both the function and the argument. The remaining cases are similar, using induction for each subterm, and extending the map whenever the construct introduces a bound variable. The induction is over the derivation that the term is well typed, so extending the context doesn't invalidate the inductive hypothesis. Equivalently, the recursion terminates because the second argument always grows smaller, even though the first argument sometimes grows larger. We have three important corollaries, each proved by constructing a suitable map between contexts. First, a closed term can be weakened to any context: ```agda weaken : ∀ {Γ M A} → ∅ ⊢ M ⦂ A ---------- → Γ ⊢ M ⦂ A weaken {Γ} ⊢M = rename ρ ⊢M where ρ : ∀ {z C} → ∅ ∋ z ⦂ C --------- → Γ ∋ z ⦂ C ρ () ``` Here the map `ρ` is trivial, since there are no possible arguments in the empty context `∅`. Second, if the last two variables in a context are equal then we can drop the shadowed one: ```agda drop : ∀ {Γ x M A B C} → Γ , x ⦂ A , x ⦂ B ⊢ M ⦂ C -------------------------- → Γ , x ⦂ B ⊢ M ⦂ C drop {Γ} {x} {M} {A} {B} {C} ⊢M = rename ρ ⊢M where ρ : ∀ {z C} → Γ , x ⦂ A , x ⦂ B ∋ z ⦂ C ------------------------- → Γ , x ⦂ B ∋ z ⦂ C ρ Z = Z ρ (S x≢x Z) = ⊥-elim (x≢x refl) ρ (S z≢x (S _ ∋z)) = S z≢x ∋z ``` Here map `ρ` can never be invoked on the inner occurrence of `x` since it is masked by the outer occurrence. Skipping over the `x` in the first position can only happen if the variable looked for differs from `x` (the evidence for which is `x≢x` or `z≢x`) but if the variable is found in the second position, which also contains `x`, this leads to a contradiction (evidenced by `x≢x refl`). Third, if the last two variables in a context differ then we can swap them: ```agda swap : ∀ {Γ x y M A B C} → x ≢ y → Γ , y ⦂ B , x ⦂ A ⊢ M ⦂ C -------------------------- → Γ , x ⦂ A , y ⦂ B ⊢ M ⦂ C swap {Γ} {x} {y} {M} {A} {B} {C} x≢y ⊢M = rename ρ ⊢M where ρ : ∀ {z C} → Γ , y ⦂ B , x ⦂ A ∋ z ⦂ C -------------------------- → Γ , x ⦂ A , y ⦂ B ∋ z ⦂ C ρ Z = S x≢y Z ρ (S z≢x Z) = Z ρ (S z≢x (S z≢y ∋z)) = S z≢y (S z≢x ∋z) ``` Here the renaming map takes a variable at the end into a variable one from the end, and vice versa. The first line is responsible for moving `x` from a position at the end to a position one from the end with `y` at the end, and requires the provided evidence that `x ≢ y`. ## Substitution The key to preservation – and the trickiest bit of the proof – is the lemma establishing that substitution preserves types. Recall that in order to avoid renaming bound variables, substitution is restricted to be by closed terms only. This restriction was not enforced by our definition of substitution, but it is captured by our lemma to assert that substitution preserves typing. Our concern is with reducing closed terms, which means that when we apply `β` reduction, the term substituted in contains a single free variable (the bound variable of the lambda abstraction, or similarly for case or fixpoint). However, substitution is defined by recursion, and as we descend into terms with bound variables the context grows. So for the induction to go through, we require an arbitrary context `Γ`, as in the statement of the lemma. Here is the formal statement and proof that substitution preserves types: ```agda subst : ∀ {Γ x N V A B} → ∅ ⊢ V ⦂ A → Γ , x ⦂ A ⊢ N ⦂ B -------------------- → Γ ⊢ N [ x := V ] ⦂ B subst {x = y} ⊢V (⊢` {x = x} Z) with x ≟ y ... | yes _ = weaken ⊢V ... | no x≢y = ⊥-elim (x≢y refl) subst {x = y} ⊢V (⊢` {x = x} (S x≢y ∋x)) with x ≟ y ... | yes refl = ⊥-elim (x≢y refl) ... | no _ = ⊢` ∋x subst {x = y} ⊢V (⊢ƛ {x = x} ⊢N) with x ≟ y ... | yes refl = ⊢ƛ (drop ⊢N) ... | no x≢y = ⊢ƛ (subst ⊢V (swap x≢y ⊢N)) subst ⊢V (⊢L · ⊢M) = (subst ⊢V ⊢L) · (subst ⊢V ⊢M) subst ⊢V ⊢zero = ⊢zero subst ⊢V (⊢suc ⊢M) = ⊢suc (subst ⊢V ⊢M) subst {x = y} ⊢V (⊢case {x = x} ⊢L ⊢M ⊢N) with x ≟ y ... | yes refl = ⊢case (subst ⊢V ⊢L) (subst ⊢V ⊢M) (drop ⊢N) ... | no x≢y = ⊢case (subst ⊢V ⊢L) (subst ⊢V ⊢M) (subst ⊢V (swap x≢y ⊢N)) subst {x = y} ⊢V (⊢μ {x = x} ⊢M) with x ≟ y ... | yes refl = ⊢μ (drop ⊢M) ... | no x≢y = ⊢μ (subst ⊢V (swap x≢y ⊢M)) ``` We induct on the evidence that `N` is well typed in the context `Γ` extended by `x`. First, we note a wee issue with naming. In the lemma statement, the variable `x` is an implicit parameter for the variable substituted, while in the type rules for variables, abstractions, cases, and fixpoints, the variable `x` is an implicit parameter for the relevant variable. We are going to need to get hold of both variables, so we use the syntax `{x = y}` to bind `y` to the substituted variable and the syntax `{x = x}` to bind `x` to the relevant variable in the patterns for `` ⊢` ``, `⊢ƛ`, `⊢case`, and `⊢μ`. Using the name `y` here is consistent with the naming in the original definition of substitution in the previous chapter. The proof never mentions the types of `x`, `y`, `V`, or `N`, so in what follows we choose type names as convenient. Now that naming is resolved, let's unpack the first three cases: * In the variable case, we must show: ∅ ⊢ V ⦂ B Γ , y ⦂ B ⊢ ` x ⦂ A ------------------------ Γ ⊢ ` x [ y := V ] ⦂ A where the second hypothesis follows from: Γ , y ⦂ B ∋ x ⦂ A There are two subcases, depending on the evidence for this judgment: + The lookup judgment is evidenced by rule `Z`: ---------------- Γ , x ⦂ A ∋ x ⦂ A In this case, `x` and `y` are necessarily identical, as are `A` and `B`. Nonetheless, we must evaluate `x ≟ y` in order to allow the definition of substitution to simplify: - If the variables are equal, then after simplification we must show: ∅ ⊢ V ⦂ A --------- Γ ⊢ V ⦂ A which follows by weakening. - If the variables are unequal we have a contradiction. + The lookup judgment is evidenced by rule `S`: x ≢ y Γ ∋ x ⦂ A ----------------- Γ , y ⦂ B ∋ x ⦂ A In this case, `x` and `y` are necessarily distinct. Nonetheless, we must again evaluate `x ≟ y` in order to allow the definition of substitution to simplify: - If the variables are equal we have a contradiction. - If the variables are unequal, then after simplification we must show: ∅ ⊢ V ⦂ B x ≢ y Γ ∋ x ⦂ A ------------- Γ ⊢ ` x ⦂ A which follows by the typing rule for variables. * In the abstraction case, we must show: ∅ ⊢ V ⦂ B Γ , y ⦂ B ⊢ (ƛ x ⇒ N) ⦂ A ⇒ C -------------------------------- Γ ⊢ (ƛ x ⇒ N) [ y := V ] ⦂ A ⇒ C where the second hypothesis follows from: Γ , y ⦂ B , x ⦂ A ⊢ N ⦂ C We evaluate `x ≟ y` in order to allow the definition of substitution to simplify: + If the variables are equal then after simplification we must show: ∅ ⊢ V ⦂ B Γ , x ⦂ B , x ⦂ A ⊢ N ⦂ C ------------------------- Γ ⊢ ƛ x ⇒ N ⦂ A ⇒ C From the drop lemma we know: Γ , x ⦂ B , x ⦂ A ⊢ N ⦂ C ------------------------- Γ , x ⦂ A ⊢ N ⦂ C The typing rule for abstractions then yields the required conclusion. + If the variables are distinct then after simplification we must show: ∅ ⊢ V ⦂ B x ≢ y Γ , y ⦂ B , x ⦂ A ⊢ N ⦂ C -------------------------------- Γ ⊢ ƛ x ⇒ (N [ y := V ]) ⦂ A ⇒ C From the swap lemma we know: x ≢ y Γ , y ⦂ B , x ⦂ A ⊢ N ⦂ C ------------------------- Γ , x ⦂ A , y ⦂ B ⊢ N ⦂ C The inductive hypothesis gives us: ∅ ⊢ V ⦂ B Γ , x ⦂ A , y ⦂ B ⊢ N ⦂ C ---------------------------- Γ , x ⦂ A ⊢ N [ y := V ] ⦂ C The typing rule for abstractions then yields the required conclusion. * In the application case, we must show: ∅ ⊢ V ⦂ C Γ , y ⦂ C ⊢ L · M ⦂ B -------------------------- Γ ⊢ (L · M) [ y := V ] ⦂ B where the second hypothesis follows from the two judgments: Γ , y ⦂ C ⊢ L ⦂ A ⇒ B Γ , y ⦂ C ⊢ M ⦂ A By the definition of substitution, we must show: ∅ ⊢ V ⦂ C Γ , y ⦂ C ⊢ L ⦂ A ⇒ B Γ , y ⦂ C ⊢ M ⦂ A --------------------------------------- Γ ⊢ (L [ y := V ]) · (M [ y := V ]) ⦂ B Applying the induction hypothesis for `L` and `M` and the typing rule for applications yields the required conclusion. The remaining cases are similar, using induction for each subterm. Where the construct introduces a bound variable we need to compare it with the substituted variable, applying the drop lemma if they are equal and the swap lemma if they are distinct. For Agda it makes a difference whether we write `x ≟ y` or `y ≟ x`. In an interactive proof, Agda will show which residual `with` clauses in the definition of `_[_:=_]` need to be simplified, and the `with` clauses in `subst` need to match these exactly. The guideline is that Agda knows nothing about symmetry or commutativity, which require invoking appropriate lemmas, so it is important to think about order of arguments and to be consistent. #### Exercise `subst′` (stretch) Rewrite `subst` to work with the modified definition `_[_:=_]′` from the exercise in the previous chapter. As before, this should factor dealing with bound variables into a single function, defined by mutual recursion with the proof that substitution preserves types. ```agda -- Definition from previous chapter's solution: mutual infix 9 _[_:=_]′ _⇒′_[_:=_]′ _[_:=_]′ : Term → Id → Term → Term (` x) [ y := V ]′ with x ≟ y ... | yes _ = V ... | no _ = ` x (ƛ x ⇒ N) [ y := V ]′ = ƛ x ⇒ (x ⇒′ N [ y := V ]′) (L · M) [ y := V ]′ = L [ y := V ]′ · M [ y := V ]′ (`zero) [ y := V ]′ = `zero (`suc M) [ y := V ]′ = `suc M [ y := V ]′ (case L [zero⇒ M |suc x ⇒ N ]) [ y := V ]′ = case L [ y := V ]′ [zero⇒ M [ y := V ]′ |suc x ⇒ (x ⇒′ N [ y := V ]′) ] (μ x ⇒ N) [ y := V ]′ = μ x ⇒ (x ⇒′ N [ y := V ]′) _⇒′_[_:=_]′ : Id → Term → Id → Term → Term x ⇒′ M [ y := V ]′ with x ≟ y ... | yes _ = M ... | no _ = M [ y := V ]′ -- Solution for this exercise mutual subst′ : ∀ {Γ x N V A B} → ∅ ⊢ V ⦂ A → Γ , x ⦂ A ⊢ N ⦂ B -------------------- → Γ ⊢ N [ x := V ]′ ⦂ B subst′ {x = y} ⊢V (⊢` {x = x} Z) with x ≟ y ... | yes _ = weaken ⊢V ... | no x≢y = ⊥-elim (x≢y refl) subst′ {x = y} ⊢V (⊢` {x = x} (S x≢y ∋x)) with x ≟ y ... | yes refl = ⊥-elim (x≢y refl) ... | no _ = ⊢` ∋x subst′ {x = y} ⊢V (⊢ƛ {x = x} ⊢N) = ⊢ƛ (subst⇒′ ⊢V ⊢N) subst′ ⊢V (⊢L · ⊢M) = (subst′ ⊢V ⊢L) · (subst′ ⊢V ⊢M) subst′ ⊢V ⊢zero = ⊢zero subst′ ⊢V (⊢suc ⊢M) = ⊢suc (subst′ ⊢V ⊢M) subst′ {x = y} ⊢V (⊢case {x = x} ⊢L ⊢M ⊢N) = ⊢case (subst′ ⊢V ⊢L) (subst′ ⊢V ⊢M) (subst⇒′ ⊢V ⊢N) subst′ {x = y} ⊢V (⊢μ {x = x} ⊢M) = ⊢μ (subst⇒′ ⊢V ⊢M) subst⇒′ : ∀ {Γ x y N V A B C} → ∅ ⊢ V ⦂ A → Γ , y ⦂ A , x ⦂ C ⊢ N ⦂ B -------------------- → Γ , x ⦂ C ⊢ x ⇒′ N [ y := V ]′ ⦂ B subst⇒′ {x = x} {y = y} ⊢V ⊢N with x ≟ y ... | yes refl = drop ⊢N ... | no x≢y = subst′ ⊢V (swap x≢y ⊢N) ``` ## Preservation Once we have shown that substitution preserves types, showing that reduction preserves types is straightforward: ```agda preserve : ∀ {M N A} → ∅ ⊢ M ⦂ A → M —→ N ---------- → ∅ ⊢ N ⦂ A preserve (⊢` ()) preserve (⊢ƛ ⊢N) () preserve (⊢L · ⊢M) (ξ-·₁ L—→L′) = (preserve ⊢L L—→L′) · ⊢M preserve (⊢L · ⊢M) (ξ-·₂ VL M—→M′) = ⊢L · (preserve ⊢M M—→M′) preserve ((⊢ƛ ⊢N) · ⊢V) (β-ƛ VV) = subst ⊢V ⊢N preserve ⊢zero () preserve (⊢suc ⊢M) (ξ-suc M—→M′) = ⊢suc (preserve ⊢M M—→M′) preserve (⊢case ⊢L ⊢M ⊢N) (ξ-case L—→L′) = ⊢case (preserve ⊢L L—→L′) ⊢M ⊢N preserve (⊢case ⊢zero ⊢M ⊢N) (β-zero) = ⊢M preserve (⊢case (⊢suc ⊢V) ⊢M ⊢N) (β-suc VV) = subst ⊢V ⊢N preserve (⊢μ ⊢M) (β-μ) = subst (⊢μ ⊢M) ⊢M ``` The proof never mentions the types of `M` or `N`, so in what follows we choose type name as convenient. Let's unpack the cases for two of the reduction rules: * Rule `ξ-·₁`. We have L —→ L′ ---------------- L · M —→ L′ · M where the left-hand side is typed by Γ ⊢ L ⦂ A ⇒ B Γ ⊢ M ⦂ A ------------- Γ ⊢ L · M ⦂ B By induction, we have Γ ⊢ L ⦂ A ⇒ B L —→ L′ -------------- Γ ⊢ L′ ⦂ A ⇒ B from which the typing of the right-hand side follows immediately. * Rule `β-ƛ`. We have Value V ----------------------------- (ƛ x ⇒ N) · V —→ N [ x := V ] where the left-hand side is typed by Γ , x ⦂ A ⊢ N ⦂ B ------------------- Γ ⊢ ƛ x ⇒ N ⦂ A ⇒ B Γ ⊢ V ⦂ A -------------------------------- Γ ⊢ (ƛ x ⇒ N) · V ⦂ B By the substitution lemma, we have Γ ⊢ V ⦂ A Γ , x ⦂ A ⊢ N ⦂ B -------------------- Γ ⊢ N [ x := V ] ⦂ B from which the typing of the right-hand side follows immediately. The remaining cases are similar. Each `ξ` rule follows by induction, and each `β` rule follows by the substitution lemma. ## Evaluation By repeated application of progress and preservation, we can evaluate any well-typed term. In this section, we will present an Agda function that computes the reduction sequence from any given closed, well-typed term to its value, if it has one. Some terms may reduce forever. Here is a simple example: ```agda sucμ = μ "x" ⇒ `suc (` "x") _ = begin sucμ —→⟨ β-μ ⟩ `suc sucμ —→⟨ ξ-suc β-μ ⟩ `suc `suc sucμ —→⟨ ξ-suc (ξ-suc β-μ) ⟩ `suc `suc `suc sucμ -- ... ∎ ``` Since every Agda computation must terminate, we cannot simply ask Agda to reduce a term to a value. Instead, we will provide a natural number to Agda, and permit it to stop short of a value if the term requires more than the given number of reduction steps. A similar issue arises with cryptocurrencies. Systems which use smart contracts require the miners that maintain the blockchain to evaluate the program which embodies the contract. For instance, validating a transaction on Ethereum may require executing a program for the Ethereum Virtual Machine (EVM). A long-running or non-terminating program might cause the miner to invest arbitrary effort in validating a contract for little or no return. To avoid this situation, each transaction is accompanied by an amount of _gas_ available for computation. Each step executed on the EVM is charged an advertised amount of gas, and the transaction pays for the gas at a published rate: a given number of Ethers (the currency of Ethereum) per unit of gas. By analogy, we will use the name _gas_ for the parameter which puts a bound on the number of reduction steps. `Gas` is specified by a natural number: ```agda record Gas : Set where constructor gas field amount : ℕ ``` When our evaluator returns a term `N`, it will either give evidence that `N` is a value or indicate that it ran out of gas: ```agda data Finished (N : Term) : Set where done : Value N ---------- → Finished N out-of-gas : ---------- Finished N ``` Given a term `L` of type `A`, the evaluator will, for some `N`, return a reduction sequence from `L` to `N` and an indication of whether reduction finished: ```agda data Steps (L : Term) : Set where steps : ∀ {N} → L —↠ N → Finished N ---------- → Steps L ``` The evaluator takes gas and evidence that a term is well typed, and returns the corresponding steps: ```agda eval : ∀ {L A} → Gas → ∅ ⊢ L ⦂ A --------- → Steps L eval {L} (gas zero) ⊢L = steps (L ∎) out-of-gas eval {L} (gas (suc m)) ⊢L with progress ⊢L ... | done VL = steps (L ∎) (done VL) ... | step {M} L—→M with eval (gas m) (preserve ⊢L L—→M) ... | steps M—↠N fin = steps (L —→⟨ L—→M ⟩ M—↠N) fin ``` Let `L` be the name of the term we are reducing, and `⊢L` be the evidence that `L` is well typed. We consider the amount of gas remaining. There are two possibilities: * It is zero, so we stop early. We return the trivial reduction sequence `L —↠ L` and an indication that we are out of gas. * It is non-zero and after the next step we have `m` gas remaining. Apply progress to the evidence that term `L` is well typed. There are two possibilities: + Term `L` is a value, so we are done. We return the trivial reduction sequence `L —↠ L` and the evidence that `L` is a value. + Term `L` steps to another term `M`. Preservation provides evidence that `M` is also well typed, and we recursively invoke `eval` on the remaining gas. The result is evidence that `M —↠ N` and indication of whether reduction finished. We combine the evidence that `L —→ M` and `M —↠ N` to return evidence that `L —↠ N` and the indication of whether reduction finished. ### Examples We can now use Agda to compute the non-terminating reduction sequence given earlier. First, we show that the term `sucμ` is well typed: ```agda ⊢sucμ : ∅ ⊢ μ "x" ⇒ `suc ` "x" ⦂ `ℕ ⊢sucμ = ⊢μ (⊢suc (⊢` ∋x)) where ∋x = Z ``` To show the first three steps of the infinite reduction sequence, we evaluate with three steps worth of gas: ```agda _ : eval (gas 3) ⊢sucμ ≡ steps (μ "x" ⇒ `suc ` "x" —→⟨ β-μ ⟩ `suc (μ "x" ⇒ `suc ` "x") —→⟨ ξ-suc β-μ ⟩ `suc (`suc (μ "x" ⇒ `suc ` "x")) —→⟨ ξ-suc (ξ-suc β-μ) ⟩ `suc (`suc (`suc (μ "x" ⇒ `suc ` "x"))) ∎) out-of-gas _ = refl ``` Similarly, we can use Agda to compute the reduction sequences given in the previous chapter. We start with the Church numeral two applied to successor and zero. Supplying 100 steps of gas is more than enough: ```agda _ : eval (gas 100) (⊢twoᶜ · ⊢sucᶜ · ⊢zero) ≡ steps ((ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero —→⟨ ξ-·₁ (β-ƛ V-ƛ) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · `zero —→⟨ β-ƛ V-zero ⟩ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · `zero) —→⟨ ξ-·₂ V-ƛ (β-ƛ V-zero) ⟩ (ƛ "n" ⇒ `suc ` "n") · `suc `zero —→⟨ β-ƛ (V-suc V-zero) ⟩ `suc (`suc `zero) ∎) (done (V-suc (V-suc V-zero))) _ = refl ``` The example above was generated by using `C-c C-n` to normalise the left-hand side of the equation and pasting in the result as the right-hand side of the equation. The example reduction of the previous chapter was derived from this result, reformatting and writing `twoᶜ` and `sucᶜ` in place of their expansions. Next, we show two plus two is four: ```agda _ : eval (gas 100) ⊢2+2 ≡ steps ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · `suc (`suc `zero) · `suc (`suc `zero) —→⟨ ξ-·₁ (ξ-·₁ β-μ) ⟩ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ])) · `suc (`suc `zero) · `suc (`suc `zero) —→⟨ ξ-·₁ (β-ƛ (V-suc (V-suc V-zero))) ⟩ (ƛ "n" ⇒ case `suc (`suc `zero) [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ]) · `suc (`suc `zero) —→⟨ β-ƛ (V-suc (V-suc V-zero)) ⟩ case `suc (`suc `zero) [zero⇒ `suc (`suc `zero) |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · `suc (`suc `zero)) ] —→⟨ β-suc (V-suc V-zero) ⟩ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · `suc `zero · `suc (`suc `zero)) —→⟨ ξ-suc (ξ-·₁ (ξ-·₁ β-μ)) ⟩ `suc ((ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ])) · `suc `zero · `suc (`suc `zero)) —→⟨ ξ-suc (ξ-·₁ (β-ƛ (V-suc V-zero))) ⟩ `suc ((ƛ "n" ⇒ case `suc `zero [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ]) · `suc (`suc `zero)) —→⟨ ξ-suc (β-ƛ (V-suc (V-suc V-zero))) ⟩ `suc case `suc `zero [zero⇒ `suc (`suc `zero) |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · `suc (`suc `zero)) ] —→⟨ ξ-suc (β-suc V-zero) ⟩ `suc (`suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · `zero · `suc (`suc `zero))) —→⟨ ξ-suc (ξ-suc (ξ-·₁ (ξ-·₁ β-μ))) ⟩ `suc (`suc ((ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ])) · `zero · `suc (`suc `zero))) —→⟨ ξ-suc (ξ-suc (ξ-·₁ (β-ƛ V-zero))) ⟩ `suc (`suc ((ƛ "n" ⇒ case `zero [zero⇒ ` "n" |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ]) · `suc (`suc `zero))) —→⟨ ξ-suc (ξ-suc (β-ƛ (V-suc (V-suc V-zero)))) ⟩ `suc (`suc case `zero [zero⇒ `suc (`suc `zero) |suc "m" ⇒ `suc ((μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "m" · `suc (`suc `zero)) ]) —→⟨ ξ-suc (ξ-suc β-zero) ⟩ `suc (`suc (`suc (`suc `zero))) ∎) (done (V-suc (V-suc (V-suc (V-suc V-zero))))) _ = refl ``` Again, the derivation in the previous chapter was derived by editing the above. Similarly, we can evaluate the corresponding term for Church numerals: ```agda _ : eval (gas 100) ⊢2+2ᶜ ≡ steps ((ƛ "m" ⇒ (ƛ "n" ⇒ (ƛ "s" ⇒ (ƛ "z" ⇒ ` "m" · ` "s" · (` "n" · ` "s" · ` "z"))))) · (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero —→⟨ ξ-·₁ (ξ-·₁ (ξ-·₁ (β-ƛ V-ƛ))) ⟩ (ƛ "n" ⇒ (ƛ "s" ⇒ (ƛ "z" ⇒ (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · ` "s" · (` "n" · ` "s" · ` "z")))) · (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero —→⟨ ξ-·₁ (ξ-·₁ (β-ƛ V-ƛ)) ⟩ (ƛ "s" ⇒ (ƛ "z" ⇒ (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · ` "s" · ((ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · ` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero —→⟨ ξ-·₁ (β-ƛ V-ƛ) ⟩ (ƛ "z" ⇒ (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · ((ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · ` "z")) · `zero —→⟨ β-ƛ V-zero ⟩ (ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · ((ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero) —→⟨ ξ-·₁ (β-ƛ V-ƛ) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · ((ƛ "s" ⇒ (ƛ "z" ⇒ ` "s" · (` "s" · ` "z"))) · (ƛ "n" ⇒ `suc ` "n") · `zero) —→⟨ ξ-·₂ V-ƛ (ξ-·₁ (β-ƛ V-ƛ)) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · ((ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · `zero) —→⟨ ξ-·₂ V-ƛ (β-ƛ V-zero) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · ((ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · `zero)) —→⟨ ξ-·₂ V-ƛ (ξ-·₂ V-ƛ (β-ƛ V-zero)) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · ((ƛ "n" ⇒ `suc ` "n") · `suc `zero) —→⟨ ξ-·₂ V-ƛ (β-ƛ (V-suc V-zero)) ⟩ (ƛ "z" ⇒ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · ` "z")) · `suc (`suc `zero) —→⟨ β-ƛ (V-suc (V-suc V-zero)) ⟩ (ƛ "n" ⇒ `suc ` "n") · ((ƛ "n" ⇒ `suc ` "n") · `suc (`suc `zero)) —→⟨ ξ-·₂ V-ƛ (β-ƛ (V-suc (V-suc V-zero))) ⟩ (ƛ "n" ⇒ `suc ` "n") · `suc (`suc (`suc `zero)) —→⟨ β-ƛ (V-suc (V-suc (V-suc V-zero))) ⟩ `suc (`suc (`suc (`suc `zero))) ∎) (done (V-suc (V-suc (V-suc (V-suc V-zero))))) _ = refl ``` And again, the example in the previous section was derived by editing the above. #### Exercise `mul-eval` (recommended) Using the evaluator, confirm that two times two is four. ```agda -- Definition from previous chapter's solution: mul : Term mul = μ "*" ⇒ ƛ "m" ⇒ ƛ "n" ⇒ case ` "m" [zero⇒ `zero |suc "m" ⇒ plus · ` "n" · (` "*" · ` "m" · ` "n") ] ⊢mul : ∅ ⊢ mul ⦂ (`ℕ ⇒ `ℕ ⇒ `ℕ) ⊢mul = ⊢μ (⊢ƛ (⊢ƛ (⊢case ⊢n ⊢zero ((⊢plus · ⊢n') · ((⊢* · ⊢m) · ⊢n'))))) where ⊢m = ⊢` Z ⊢n = ⊢` (S (λ()) Z) ⊢n' = ⊢` (S (λ()) Z) ⊢* = ⊢` (S (λ()) (S (λ()) (S (λ()) Z))) -- Solution for this exercise: -- Given `eval`, Agda can compute the proof term for us. -- Just place the cursor in the hole below and hit -- `Ctrl-C Ctrl-N`. This causes Agda to normalize (= evaluate) the term in -- the hole and display the result, which can then be copy pasted in the hole, -- as shown below in `mul-eval'`. mul-eval : eval (gas 100) ⊢mul ≡ {!eval (gas 100) ⊢mul!} mul-eval = refl mul-eval' : eval (gas 100) ⊢mul ≡ steps (μ "*" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ `zero |suc "m" ⇒ (μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "n" · (` "*" · ` "m" · ` "n") ])) —→⟨ β-μ ⟩ ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ `zero |suc "m" ⇒ (μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "n" · ((μ "*" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ `zero |suc "m" ⇒ (μ "+" ⇒ (ƛ "m" ⇒ (ƛ "n" ⇒ case ` "m" [zero⇒ ` "n" |suc "m" ⇒ `suc (` "+" · ` "m" · ` "n") ]))) · ` "n" · (` "*" · ` "m" · ` "n") ]))) · ` "m" · ` "n") ]) ∎) (done V-ƛ) mul-eval' = refl ``` #### Exercise: `progress-preservation` (practice) Without peeking at their statements above, write down the progress and preservation theorems for the simply typed lambda-calculus. ```agda -- Solution is already in the book. -- This is a pretty good exercise. -- I highly recommend to do it! ``` #### Exercise `subject_expansion` (practice) We say that `M` _reduces_ to `N` if `M —→ N`, but we can also describe the same situation by saying that `N` _expands_ to `M`. The preservation property is sometimes called _subject reduction_. Its opposite is _subject expansion_, which holds if `M —→ N` and `∅ ⊢ N ⦂ A` imply `∅ ⊢ M ⦂ A`. Find two counter-examples to subject expansion, one with case expressions and one not involving case expressions. ```agda -- The following term is not well-typed, because the case-branch for `suc` -- uses the number `zero as a function. -- -- However since the `case`-expression is applied to `zero`, the -- ill-typed branch is just thrown away and the whole expression -- reduces to `zero`. counter-example₁ : Term counter-example₁ = case `zero [zero⇒ `zero |suc "x" ⇒ (`zero · `zero) ] -- A similar counter-example can be constructed by using two -- well-typed terms in the branches of a case, but both terms have -- *different* types, which is not allowed by the ⊢case typing rule. counter-example₂ : Term counter-example₂ = case `zero [zero⇒ `zero |suc "x" ⇒ (ƛ "x" ⇒ ` "x") ] ``` ## Well-typed terms don't get stuck A term is _normal_ if it cannot reduce: ```agda Normal : Term → Set Normal M = ∀ {N} → ¬ (M —→ N) ``` A term is _stuck_ if it is normal yet not a value: ```agda Stuck : Term → Set Stuck M = Normal M × ¬ Value M ``` Using progress, it is easy to show that no well-typed term is stuck: ```agda -- postulate -- unstuck : ∀ {M A} -- → ∅ ⊢ M ⦂ A -- ----------- -- → ¬ (Stuck M) ``` Using preservation, it is easy to show that after any number of steps, a well-typed term remains well typed: ```agda -- postulate -- preserves : ∀ {M N A} -- → ∅ ⊢ M ⦂ A -- → M —↠ N -- --------- -- → ∅ ⊢ N ⦂ A ``` An easy consequence is that starting from a well-typed term, taking any number of reduction steps leads to a term that is not stuck: ```agda -- postulate -- wttdgs : ∀ {M N A} -- → ∅ ⊢ M ⦂ A -- → M —↠ N -- ----------- -- → ¬ (Stuck N) ``` Felleisen and Wright, who introduced proofs via progress and preservation, summarised this result with the slogan _well-typed terms don't get stuck_. (They were referring to earlier work by Robin Milner, who used denotational rather than operational semantics. He introduced `wrong` as the denotation of a term with a type error, and showed _well-typed terms don't go wrong_.) #### Exercise `stuck` (practice) Give an example of an ill-typed term that does get stuck. ```agda -- The following term is ill-typed because it uses a number as a function. -- It is neither a value, nor does any of the reduction rules for _·_ apply. stuck-example : Term stuck-example = `zero · `zero ``` #### Exercise `unstuck` (recommended) Provide proofs of the three postulates, `unstuck`, `preserves`, and `wttdgs` above. ```agda unstuck : ∀ {M A} → ∅ ⊢ M ⦂ A ----------- → ¬ (Stuck M) unstuck ⊢M ⟨ normal-M , ¬val-M ⟩ with progress ⊢M ... | step M–→M' = normal-M M–→M' ... | done val-M = ¬val-M val-M preserves : ∀ {M N A} → ∅ ⊢ M ⦂ A → M —↠ N --------- → ∅ ⊢ N ⦂ A preserves ⊢M (_ ∎) = ⊢M preserves ⊢M (_ —→⟨ M—→N ⟩ N—↠P) = preserves (preserve ⊢M M—→N) N—↠P wttdgs : ∀ {M N A} → ∅ ⊢ M ⦂ A → M —↠ N ----------- → ¬ (Stuck N) wttdgs ⊢M M—↠N = unstuck (preserves ⊢M M—↠N) ``` ## Reduction is deterministic When we introduced reduction, we claimed it was deterministic. For completeness, we present a formal proof here. Our proof will need a variant of congruence to deal with functions of four arguments (to deal with `case_[zero⇒_|suc_⇒_]`). It is exactly analogous to `cong` and `cong₂` as defined previously: ```agda cong₄ : ∀ {A B C D E : Set} (f : A → B → C → D → E) {s w : A} {t x : B} {u y : C} {v z : D} → s ≡ w → t ≡ x → u ≡ y → v ≡ z → f s t u v ≡ f w x y z cong₄ f refl refl refl refl = refl ``` It is now straightforward to show that reduction is deterministic: ```agda det : ∀ {M M′ M″} → (M —→ M′) → (M —→ M″) -------- → M′ ≡ M″ det (ξ-·₁ L—→L′) (ξ-·₁ L—→L″) = cong₂ _·_ (det L—→L′ L—→L″) refl det (ξ-·₁ L—→L′) (ξ-·₂ VL M—→M″) = ⊥-elim (V¬—→ VL L—→L′) det (ξ-·₁ L—→L′) (β-ƛ _) = ⊥-elim (V¬—→ V-ƛ L—→L′) det (ξ-·₂ VL _) (ξ-·₁ L—→L″) = ⊥-elim (V¬—→ VL L—→L″) det (ξ-·₂ _ M—→M′) (ξ-·₂ _ M—→M″) = cong₂ _·_ refl (det M—→M′ M—→M″) det (ξ-·₂ _ M—→M′) (β-ƛ VM) = ⊥-elim (V¬—→ VM M—→M′) det (β-ƛ _) (ξ-·₁ L—→L″) = ⊥-elim (V¬—→ V-ƛ L—→L″) det (β-ƛ VM) (ξ-·₂ _ M—→M″) = ⊥-elim (V¬—→ VM M—→M″) det (β-ƛ _) (β-ƛ _) = refl det (ξ-suc M—→M′) (ξ-suc M—→M″) = cong `suc_ (det M—→M′ M—→M″) det (ξ-case L—→L′) (ξ-case L—→L″) = cong₄ case_[zero⇒_|suc_⇒_] (det L—→L′ L—→L″) refl refl refl det (ξ-case L—→L′) β-zero = ⊥-elim (V¬—→ V-zero L—→L′) det (ξ-case L—→L′) (β-suc VL) = ⊥-elim (V¬—→ (V-suc VL) L—→L′) det β-zero (ξ-case M—→M″) = ⊥-elim (V¬—→ V-zero M—→M″) det β-zero β-zero = refl det (β-suc VL) (ξ-case L—→L″) = ⊥-elim (V¬—→ (V-suc VL) L—→L″) det (β-suc _) (β-suc _) = refl det β-μ β-μ = refl ``` The proof is by induction over possible reductions. We consider three typical cases: * Two instances of `ξ-·₁`: L —→ L′ L —→ L″ --------------- ξ-·₁ --------------- ξ-·₁ L · M —→ L′ · M L · M —→ L″ · M By induction we have `L′ ≡ L″`, and hence by congruence `L′ · M ≡ L″ · M`. * An instance of `ξ-·₁` and an instance of `ξ-·₂`: Value L L —→ L′ M —→ M″ --------------- ξ-·₁ --------------- ξ-·₂ L · M —→ L′ · M L · M —→ L · M″ The rule on the left requires `L` to reduce, but the rule on the right requires `L` to be a value. This is a contradiction since values do not reduce. If the value constraint was removed from `ξ-·₂`, or from one of the other reduction rules, then determinism would no longer hold. * Two instances of `β-ƛ`: Value V Value V ----------------------------- β-ƛ ----------------------------- β-ƛ (ƛ x ⇒ N) · V —→ N [ x := V ] (ƛ x ⇒ N) · V —→ N [ x := V ] Since the left-hand sides are identical, the right-hand sides are also identical. The formal proof simply invokes `refl`. Five of the 18 lines in the above proof are redundant, e.g., the case when one rule is `ξ-·₁` and the other is `ξ-·₂` is considered twice, once with `ξ-·₁` first and `ξ-·₂` second, and the other time with the two swapped. What we might like to do is delete the redundant lines and add det M—→M′ M—→M″ = sym (det M—→M″ M—→M′) to the bottom of the proof. But this does not work: the termination checker complains, because the arguments have merely switched order and neither is smaller. #### Quiz Suppose we add a new term `zap` with the following reduction rule -------- β-zap M —→ zap and the following typing rule: ----------- ⊢zap Γ ⊢ zap ⦂ A Which of the following properties remain true in the presence of these rules? For each property, write either "remains true" or "becomes false." If a property becomes false, give a counterexample: - Determinism - Progress - Preservation ``` -- Answer: -- -- - Determinism does not hold anymore, because each term can also reduce to `zap`. -- -- - Progress does still hold, because `zap` terms can reduce infinitely often -- by using `β-zap` with `M = zap`, which allows to reduce `zap –→ zap`. -- -- - Preservation does still hold, because a `zap` term can assume any type. -- This means that in whatever context a `zap` term appears, it can -- just assume the type the term had before reducing it to `zap`. ``` #### Quiz Suppose instead that we add a new term `foo` with the following reduction rules: ------------------ β-foo₁ (λ x ⇒ ` x) —→ foo ----------- β-foo₂ foo —→ zero Which of the following properties remain true in the presence of this rule? For each one, write either "remains true" or else "becomes false." If a property becomes false, give a counterexample: - Determinism - Progress - Preservation ``` -- Answer: -- -- - Determinism does not hold anymore: -- The term `(λ x ⇒ ` x) · `zero` can reduce in two ways: -- 1. to `zero` by performing the function application. -- 2. to `foo · `zero` by using β-foo₁. -- -- - Progress does still hold: Any term which reduced before, can still -- reduce in the same way, and because we didn't add any typing rule for `foo`, -- we don't need to consider new terms in progress (which assumes a well-typed term). -- -- - Preservation does not hold anymore, due to the counter-example for determinism: -- `(λ x ⇒ ` x)` is a well-typed term, but reduces to `foo`, for -- which we don't have any typing rules. -- Even if there were typing rules for foo, preservation would still fail, because -- reducing two steps via -- (λ x ⇒ ` x) –→ foo –→ `zero -- clearly does not preserve the type of term (changes from function to natural number). ``` #### Quiz Suppose instead that we remove the rule `ξ·₁` from the step relation. Which of the following properties remain true in the absence of this rule? For each one, write either "remains true" or else "becomes false." If a property becomes false, give a counterexample: - Determinism - Progress - Preservation ``` -- Answer: -- -- - Determinism remains true. By removing a rule, we only make it -- easier to prove determinism. -- -- - Progress does not hold anymore: The following term -- ((λ x ⇒ (λ y ⇒ ` y)) · `zero) · `zero -- is well-typed, but without ξ·₁ does not reduce. -- However, with ξ·₁, the term would reduce as followed: -- ((λ x ⇒ (λ y ⇒ ` y)) · `zero) · `zero –→ -- (λ y ⇒ ` y) · `zero –→ -- `zero -- -- - Preservation does still hold. By removing a reduction rule, we -- simply need to show less cases in preservation. ``` #### Quiz We can enumerate all the computable function from naturals to naturals, by writing out all programs of type `` `ℕ ⇒ `ℕ`` in lexical order. Write `fᵢ` for the `i`'th function in this list. Say we add a typing rule that applies the above enumeration to interpret a natural as a function from naturals to naturals: Γ ⊢ L ⦂ `ℕ Γ ⊢ M ⦂ `ℕ -------------- _·ℕ_ Γ ⊢ L · M ⦂ `ℕ And that we add the corresponding reduction rule: fᵢ(m) —→ n ---------- δ i · m —→ n Which of the following properties remain true in the presence of these rules? For each one, write either "remains true" or else "becomes false." If a property becomes false, give a counterexample: - Determinism - Progress - Preservation ``` -- Answer: -- -- - Determinism remains true, because numbers in function -- position did not reduce at all before adding this rule. -- -- - Progress remains true: We can now type applications of numbers to numbers, -- but we also added a rule to reduce them. -- -- - Preservation remains true: The application -- of a number to a number as type ℕ due to the new typing rule, -- and can only reduce due to the new typing rule, which reduces it -- to a number which also has type ℕ. ``` Are all properties preserved in this case? Are there any other alterations we would wish to make to the system? ## Unicode This chapter uses the following unicode: ƛ U+019B LATIN SMALL LETTER LAMBDA WITH STROKE (\Gl-) Δ U+0394 GREEK CAPITAL LETTER DELTA (\GD or \Delta) β U+03B2 GREEK SMALL LETTER BETA (\Gb or \beta) δ U+03B4 GREEK SMALL LETTER DELTA (\Gd or \delta) μ U+03BC GREEK SMALL LETTER MU (\Gm or \mu) ξ U+03BE GREEK SMALL LETTER XI (\Gx or \xi) ρ U+03B4 GREEK SMALL LETTER RHO (\Gr or \rho) ᵢ U+1D62 LATIN SUBSCRIPT SMALL LETTER I (\_i) ᶜ U+1D9C MODIFIER LETTER SMALL C (\^c) – U+2013 EM DASH (\em) ₄ U+2084 SUBSCRIPT FOUR (\_4) ↠ U+21A0 RIGHTWARDS TWO HEADED ARROW (\rr-) ⇒ U+21D2 RIGHTWARDS DOUBLE ARROW (\=>) ∅ U+2205 EMPTY SET (\0) ∋ U+220B CONTAINS AS MEMBER (\ni) ≟ U+225F QUESTIONED EQUAL TO (\?=) ⊢ U+22A2 RIGHT TACK (\vdash or \|-) ⦂ U+2982 Z NOTATION TYPE COLON (\:)