Started rewriting with Reasoning notation

This commit is contained in:
Wen Kokke 2017-06-27 15:47:46 +01:00
commit ba15881340
No known key found for this signature in database
GPG key ID: B987B93AC7378292
6 changed files with 10161 additions and 11096 deletions

File diff suppressed because it is too large Load diff

File diff suppressed because it is too large Load diff

File diff suppressed because it is too large Load diff

View file

@ -86,8 +86,6 @@ y = id "y"
z = id "z" z = id "z"
\end{code} \end{code}
## Extensionality
## Total Maps ## Total Maps
Our main job in this chapter will be to build a definition of Our main job in this chapter will be to build a definition of
@ -112,8 +110,8 @@ TotalMap : Set → Set
TotalMap A = Id → A TotalMap A = Id → A
\end{code} \end{code}
Intuitively, a total map over anfi element type $$A$$ _is_ just a Intuitively, a total map over anfi element type `A` _is_ just a
function that can be used to look up ids, yielding $$A$$s. function that can be used to look up ids, yielding `A`s.
\begin{code} \begin{code}
module TotalMap where module TotalMap where
@ -129,10 +127,12 @@ applied to any id.
\end{code} \end{code}
More interesting is the update function, which (as before) takes 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 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. takes `x` to `v` and takes every other key to whatever `ρ` does.
\begin{code} \begin{code}
infixl 100 _,_↦_
_,_↦_ : ∀ {A} → TotalMap A → Id → A → TotalMap A _,_↦_ : ∀ {A} → TotalMap A → Id → A → TotalMap A
(ρ , x ↦ v) y with x ≟ y (ρ , x ↦ v) y with x ≟ y
... | yes x=y = v ... | yes x=y = v
@ -140,30 +140,11 @@ takes $$x$$ to $$v$$ and takes every other key to whatever $$ρ$$ does.
\end{code} \end{code}
This definition is a nice example of higher-order programming. This definition is a nice example of higher-order programming.
The update function takes a _function_ $$ρ$$ and yields a new The update function takes a _function_ `ρ` and yields a new
function that behaves like the desired map. function that behaves like the desired map.
We define handy abbreviations for updating a map two, three, or four times. 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:
<div class="note hidden">
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`.
</div>
\begin{code}
_,_↦_,_↦_ : ∀ {A} → TotalMap A → Id → A → Id → A → TotalMap A
ρ , x₁ ↦ v₁ , x₂ ↦ v₂ = (ρ , x₁ ↦ v₁), x₂ ↦ v₂
_,_↦_,_↦_,_↦_ : ∀ {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₄
\end{code}
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:
\begin{code} \begin{code}
ρ₀ : TotalMap ρ₀ : TotalMap
@ -206,9 +187,9 @@ The `always` map returns its default element for all keys:
</div> </div>
#### Exercise: 2 stars, optional (update-eq) #### Exercise: 2 stars, optional (update-eq)
Next, if we update a map $$ρ$$ at a key $$x$$ with a new value $$v$$ Next, if we update a map `ρ` at a key `x` with a new value `v`
and then look up $$x$$ in the map resulting from the update, we get and then look up `x` in the map resulting from the update, we get
back $$v$$: back `v`:
\begin{code} \begin{code}
postulate postulate
@ -227,9 +208,9 @@ back $$v$$:
</div> </div>
#### Exercise: 2 stars, optional (update-neq) #### Exercise: 2 stars, optional (update-neq)
On the other hand, if we update a map $$m$$ at a key $$x$$ and On the other hand, if we update a map `m` at a key `x` and
then look up a _different_ key $$y$$ in the resulting map, we get then look up a _different_ key `y` in the resulting map, we get
the same result that $$m$$ would have given: the same result that `m` would have given:
\begin{code} \begin{code}
update-neq : ∀ {A} (ρ : TotalMap A) (x : Id) (v : A) (y : Id) update-neq : ∀ {A} (ρ : TotalMap A) (x : Id) (v : A) (y : Id)
@ -248,11 +229,11 @@ show two maps equal we will need to postulate extensionality.
\end{code} \end{code}
#### Exercise: 2 stars, optional (update-shadow) #### Exercise: 2 stars, optional (update-shadow)
If we update a map $$ρ$$ at a key $$x$$ with a value $$v$$ and then 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 update again with the same key `x` and another value `w`, the
resulting map behaves the same (gives the same result when applied resulting map behaves the same (gives the same result when applied
to any key) as the simpler map obtained by performing just to any key) as the simpler map obtained by performing just
the second update on $$ρ$$: the second update on `ρ`:
\begin{code} \begin{code}
postulate postulate
@ -274,9 +255,9 @@ the second update on $$ρ$$:
</div> </div>
#### Exercise: 2 stars (update-same) #### Exercise: 2 stars (update-same)
Prove the following theorem, which states that if we update a map $$ρ$$ to Prove the following theorem, which states that if we update a map `ρ` to
assign key $$x$$ the same value as it already has in $$ρ$$, then the assign key `x` the same value as it already has in `ρ`, then the
result is equal to $$ρ$$: result is equal to `ρ`:
\begin{code} \begin{code}
postulate postulate
@ -297,7 +278,7 @@ result is equal to $$ρ$$:
#### Exercise: 3 stars, recommended (update-permute) #### Exercise: 3 stars, recommended (update-permute)
Prove one final property of the `update` function: If we update a map Prove one final property of the `update` function: If we update a map
$$m$$ at two distinct keys, it doesn't matter in which order we do the `m` at two distinct keys, it doesn't matter in which order we do the
updates. updates.
\begin{code} \begin{code}
@ -354,21 +335,11 @@ module PartialMap where
\end{code} \end{code}
\begin{code} \begin{code}
infixl 100 _,_↦_
_,_↦_ : ∀ {A} (ρ : PartialMap A) (x : Id) (v : A) → PartialMap A _,_↦_ : ∀ {A} (ρ : PartialMap A) (x : Id) (v : A) → PartialMap A
ρ , x ↦ v = TotalMap._,_↦_ ρ x (just v) ρ , x ↦ v = TotalMap._,_↦_ ρ x (just v)
\end{code} \end{code}
As before, we define handy abbreviations for updating a map two, three, or four times.
\begin{code}
_,_↦_,_↦_ : ∀ {A} → PartialMap A → 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₄
\end{code}
We now lift all of the basic lemmas about total maps to partial maps. We now lift all of the basic lemmas about total maps to partial maps.

View file

@ -199,6 +199,8 @@ example₁ = ⟨ step₀ ⟩ >> ⟨ step₁ ⟩ >> ⟨ step₂ ⟩ >> ⟨ step
Context : Set Context : Set
Context = PartialMap Type Context = PartialMap Type
infix 50 _⊢_∈_
data _⊢_∈_ : Context → Term → Type → Set where data _⊢_∈_ : Context → Term → Type → Set where
Ax : ∀ {Γ x A} → Ax : ∀ {Γ x A} →
Γ x ≡ just A → Γ x ≡ just A →

View file

@ -28,8 +28,8 @@ theorem.
As we saw for the simple calculus in the [Stlc]({{ "Stlc" | relative_url }}) As we saw for the simple calculus in the [Stlc]({{ "Stlc" | relative_url }})
chapter, the first step in establishing basic properties of reduction and types chapter, the first step in establishing basic properties of reduction and types
is to identify the possible _canonical forms_ (i.e., well-typed closed values) 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 belonging to each type. For `bool`, these are the boolean values `true` and
$$false$$. For arrow types, the canonical forms are lambda-abstractions. `false`. For arrow types, the canonical forms are lambda-abstractions.
\begin{code} \begin{code}
data canonical_for_ : Term → Type → Set where data canonical_for_ : Term → Type → Set where
@ -63,43 +63,43 @@ first, then the formal version.
progress : ∀ {M A} → ∅ ⊢ M ∈ A → value M ⊎ ∃ λ N → M ⟹ N progress : ∀ {M A} → ∅ ⊢ M ∈ A → value M ⊎ ∃ λ N → M ⟹ N
\end{code} \end{code}
_Proof_: By induction on the derivation of $$\vdash t : A$$. _Proof_: By induction on the derivation of `\vdash t : A`.
- The last rule of the derivation cannot be `var`, - The last rule of the derivation cannot be `var`,
since a variable is never well typed in an empty context. since a variable is never well typed in an empty context.
- The `true`, `false`, and `abs` cases are trivial, since in - The `true`, `false`, and `abs` cases are trivial, since in
each of these cases we can see by inspecting the rule that $$t$$ each of these cases we can see by inspecting the rule that `t`
is a value. is a value.
- If the last rule of the derivation is `app`, then $$t$$ has the - If the last rule of the derivation is `app`, then `t` has the
form $$t_1\;t_2$$ for som e$$t_1$$ and $$t_2$$, where we know that form `t_1\;t_2` for som e`t_1` and `t_2`, where we know that
$$t_1$$ and $$t_2$$ are also well typed in the empty context; in particular, `t_1` and `t_2` are also well typed in the empty context; in particular,
there exists a type $$B$$ such that $$\vdash t_1 : A\to T$$ and there exists a type `B` such that `\vdash t_1 : A\to T` and
$$\vdash t_2 : B$$. By the induction hypothesis, either $$t_1$$ is a `\vdash t_2 : B`. By the induction hypothesis, either `t_1` is a
value or it can take a reduction step. value or it can take a reduction step.
- If $$t_1$$ is a value, then consider $$t_2$$, which by the other - If `t_1` is a value, then consider `t_2`, which by the other
induction hypothesis must also either be a value or take a step. induction hypothesis must also either be a value or take a step.
- Suppose $$t_2$$ is a value. Since $$t_1$$ is a value with an - Suppose `t_2` is a value. Since `t_1` is a value with an
arrow type, it must be a lambda abstraction; hence $$t_1\;t_2$$ arrow type, it must be a lambda abstraction; hence `t_1\;t_2`
can take a step by `red`. can take a step by `red`.
- Otherwise, $$t_2$$ can take a step, and hence so can $$t_1\;t_2$$ - Otherwise, `t_2` can take a step, and hence so can `t_1\;t_2`
by `app2`. by `app2`.
- If $$t_1$$ can take a step, then so can $$t_1 t_2$$ by `app1`. - If `t_1` can take a step, then so can `t_1 t_2` by `app1`.
- If the last rule of the derivation is `if`, then $$t = \text{if }t_1 - If the last rule of the derivation is `if`, then `t = \text{if }t_1
\text{ then }t_2\text{ else }t_3$$, where $$t_1$$ has type $$bool$$. By \text{ then }t_2\text{ else }t_3`, where `t_1` has type `bool`. By
the IH, $$t_1$$ either is a value or takes a step. the IH, `t_1` either is a value or takes a step.
- If $$t_1$$ is a value, then since it has type $$bool$$ it must be - If `t_1` is a value, then since it has type `bool` it must be
either $$true$$ or $$false$$. If it is $$true$$, then $$t$$ steps either `true` or `false`. If it is `true`, then `t` steps
to $$t_2$$; otherwise it steps to $$t_3$$. to `t_2`; otherwise it steps to `t_3`.
- Otherwise, $$t_1$$ takes a step, and therefore so does $$t$$ (by `if`). - Otherwise, `t_1` takes a step, and therefore so does `t` (by `if`).
\begin{code} \begin{code}
progress (Ax ()) progress (Ax ())
@ -141,23 +141,23 @@ interesting proofs), the story goes like this:
- The _preservation theorem_ is proved by induction on a typing - The _preservation theorem_ is proved by induction on a typing
derivation, pretty much as we did in the [Stlc]({{ "Stlc" | relative_url }}) derivation, pretty much as we did in the [Stlc]({{ "Stlc" | relative_url }})
chapter. The one case that is significantly different is the one for the chapter. The one case that is significantly different is the one for the
$$red$$ rule, whose definition uses the substitution operation. To see that `red` rule, whose definition uses the substitution operation. To see that
this step preserves typing, we need to know that the substitution itself this step preserves typing, we need to know that the substitution itself
does. So we prove a... does. So we prove a...
- _substitution lemma_, stating that substituting a (closed) - _substitution lemma_, stating that substituting a (closed)
term $$s$$ for a variable $$x$$ in a term $$t$$ preserves the type term `s` for a variable `x` in a term `t` preserves the type
of $$t$$. The proof goes by induction on the form of $$t$$ and of `t`. The proof goes by induction on the form of `t` and
requires looking at all the different cases in the definition requires looking at all the different cases in the definition
of substitition. This time, the tricky cases are the ones for of substitition. This time, the tricky cases are the ones for
variables and for function abstractions. In both cases, we variables and for function abstractions. In both cases, we
discover that we need to take a term $$s$$ that has been shown discover that we need to take a term `s` that has been shown
to be well-typed in some context $$\Gamma$$ and consider the same to be well-typed in some context `\Gamma` and consider the same
term $$s$$ in a slightly different context $$\Gamma'$$. For this term `s` in a slightly different context `\Gamma'`. For this
we prove a... we prove a...
- _context invariance_ lemma, showing that typing is preserved - _context invariance_ lemma, showing that typing is preserved
under "inessential changes" to the context $$\Gamma$$---in under "inessential changes" to the context `\Gamma`---in
particular, changes that do not affect any of the free particular, changes that do not affect any of the free
variables of the term. And finally, for this, we need a variables of the term. And finally, for this, we need a
careful definition of... careful definition of...
@ -172,13 +172,13 @@ order...
### Free Occurrences ### Free Occurrences
A variable $$x$$ _appears free in_ a term $$M$$ if $$M$$ contains some A variable `x` _appears free in_ a term `M` if `M` contains some
occurrence of $$x$$ that is not under an abstraction over $$x$$. occurrence of `x` that is not under an abstraction over `x`.
For example: For example:
- $$y$$ appears free, but $$x$$ does not, in $$λᵀ x ∈ (A ⇒ B) ⇒ 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$$ - 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)$$ - no variables appear free in `λᵀ x ∈ (A ⇒ B) ⇒ (λᵀ y ∈ A ⇒ x ·ᵀ y)`
Formally: Formally:
@ -211,42 +211,42 @@ are really the crux of the lambda-calculus.)
### Substitution ### Substitution
To prove that substitution preserves typing, we first need a To prove that substitution preserves typing, we first need a
technical lemma connecting free variables and typing contexts: If technical lemma connecting free variables and typing contexts: If
a variable $$x$$ appears free in a term $$M$$, and if we know $$M$$ is 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 well typed in context `Γ`, then it must be the case that
$$Γ$$ assigns a type to $$x$$. `Γ` assigns a type to `x`.
\begin{code} \begin{code}
freeLemma : ∀ {x M A Γ} → x FreeIn M → Γ ⊢ M ∈ A → ∃ λ B → Γ x ≡ just B freeLemma : ∀ {x M A Γ} → x FreeIn M → Γ ⊢ M ∈ A → ∃ λ B → Γ x ≡ just B
\end{code} \end{code}
_Proof_: We show, by induction on the proof that $$x$$ appears _Proof_: We show, by induction on the proof that `x` appears
free in $$P$$, that, for all contexts $$Γ$$, if $$P$$ is well free in `P`, that, for all contexts `Γ`, if `P` is well
typed under $$Γ$$, then $$Γ$$ assigns some type to $$x$$. typed under `Γ`, then `Γ` assigns some type to `x`.
- If the last rule used was `free-varᵀ`, then $$P = x$$, and from - If the last rule used was `free-varᵀ`, then `P = x`, and from
the assumption that $$M$$ is well typed under $$Γ$$ we have the assumption that `M` is well typed under `Γ` we have
immediately that $$Γ$$ assigns a type to $$x$$. immediately that `Γ` assigns a type to `x`.
- If the last rule used was `free-·₁`, then $$P = L ·ᵀ M$$ and $$x$$ - If the last rule used was `free-·₁`, then `P = L ·ᵀ M` and `x`
appears free in $$L$$. Since $$L$$ is well typed under $$\Gamma$$, appears free in `L`. Since `L` is well typed under `\Gamma`,
we can see from the typing rules that $$L$$ must also be, and we can see from the typing rules that `L` must also be, and
the IH then tells us that $$Γ$$ assigns $$x$$ a type. the IH then tells us that `Γ` assigns `x` a type.
- Almost all the other cases are similar: $$x$$ appears free in a - Almost all the other cases are similar: `x` appears free in a
subterm of $$P$$, and since $$P$$ is well typed under $$Γ$$, we subterm of `P`, and since `P` is well typed under `Γ`, we
know the subterm of $$M$$ in which $$x$$ appears is well typed know the subterm of `M` in which `x` appears is well typed
under $$Γ$$ as well, and the IH gives us exactly the under `Γ` as well, and the IH gives us exactly the
conclusion we want. conclusion we want.
- The only remaining case is `free-λᵀ`. In this case $$P = - The only remaining case is `free-λᵀ`. In this case `P =
λᵀ y ∈ A ⇒ N$$, and $$x$$ appears free in $$N$$; we also know that λᵀ y ∈ A ⇒ N`, and `x` appears free in `N`; we also know that
$$x$$ is different from $$y$$. The difference from the previous `x` is different from `y`. The difference from the previous
cases is that whereas $$P$$ is well typed under $$\Gamma$$, its cases is that whereas `P` is well typed under `\Gamma`, its
body $$N$$ is well typed under $$(Γ , y ↦ A)$$, so the IH body `N` is well typed under `(Γ , y ↦ A)`, so the IH
allows us to conclude that $$x$$ is assigned some type by the allows us to conclude that `x` is assigned some type by the
extended context $$(Γ , y ↦ A)$$. To conclude that $$Γ$$ extended context `(Γ , y ↦ A)`. To conclude that `Γ`
assigns a type to $$x$$, we appeal the decidable equality for names assigns a type to `x`, we appeal the decidable equality for names
`_≟_`, noting that $$x$$ and $$y$$ are different variables. `_≟_`, noting that `x` and `y` are different variables.
\begin{code} \begin{code}
freeLemma free-varᵀ (Ax Γx≡justA) = (_ , Γx≡justA) freeLemma free-varᵀ (Ax Γx≡justA) = (_ , Γx≡justA)
@ -261,11 +261,11 @@ freeLemma (free-λᵀ {x} {y} y≢x x∈N) (⇒-I ⊢N) with freeLemma x∈N ⊢
... | no _ = Γx=justC ... | no _ = Γx=justC
\end{code} \end{code}
[A subtle point: if the first argument of $$free-λᵀ$$ was of type [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 `x ≢ y` rather than of type `y ≢ x`, then the type of the
term $$Γx=justC$$ would not simplify properly.] term `Γx=justC` would not simplify properly.]
Next, we'll need the fact that any term $$M$$ which is well typed in 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). the empty context is closed (it has no free variables).
#### Exercise: 2 stars, optional (∅⊢-closed) #### Exercise: 2 stars, optional (∅⊢-closed)
@ -286,11 +286,11 @@ contradiction ()
\end{code} \end{code}
</div> </div>
Sometimes, when we have a proof $$Γ ⊢ M ∈ A$$, we will need to Sometimes, when we have a proof `Γ ⊢ M ∈ A`, we will need to
replace $$Γ$$ by a different context $$Γ′$$. When is it safe replace `Γ` by a different context `Γ′`. When is it safe
to do this? Intuitively, it must at least be the case that to do this? Intuitively, it must at least be the case that
$$Γ′$$ assigns the same types as $$Γ$$ to all the variables `Γ′` assigns the same types as `Γ` to all the variables
that appear free in $$M$$. In fact, this is the only condition that that appear free in `M`. In fact, this is the only condition that
is needed. is needed.
\begin{code} \begin{code}
@ -301,45 +301,45 @@ weaken : ∀ {Γ Γ′ M A}
\end{code} \end{code}
_Proof_: By induction on the derivation of _Proof_: By induction on the derivation of
$$Γ ⊢ M ∈ A$$. `Γ ⊢ M ∈ A`.
- If the last rule in the derivation was `var`, then $$t = x$$ - If the last rule in the derivation was `var`, then `t = x`
and $$\Gamma x = T$$. By assumption, $$\Gamma' x = T$$ as well, and and `\Gamma x = T`. By assumption, `\Gamma' x = T` as well, and
hence $$\Gamma' \vdash t : T$$ by `var`. hence `\Gamma' \vdash t : T` by `var`.
- If the last rule was `abs`, then $$t = \lambda y:A. t'$$, with - If the last rule was `abs`, then `t = \lambda y:A. t'`, with
$$T = A\to B$$ and $$\Gamma, y : A \vdash t' : B$$. The `T = A\to B` and `\Gamma, y : A \vdash t' : B`. The
induction hypothesis is that, for any context $$\Gamma''$$, if induction hypothesis is that, for any context `\Gamma''`, if
$$\Gamma, y:A$$ and $$\Gamma''$$ assign the same types to all the `\Gamma, y:A` and `\Gamma''` assign the same types to all the
free variables in $$t'$$, then $$t'$$ has type $$B$$ under free variables in `t'`, then `t'` has type `B` under
$$\Gamma''$$. Let $$\Gamma'$$ be a context which agrees with `\Gamma''`. Let `\Gamma'` be a context which agrees with
$$\Gamma$$ on the free variables in $$t$$; we must show `\Gamma` on the free variables in `t`; we must show
$$\Gamma' \vdash \lambda y:A. t' : A\to B$$. `\Gamma' \vdash \lambda y:A. t' : A\to B`.
By $$abs$$, it suffices to show that $$\Gamma', y:A \vdash t' : t'$$. By `abs`, it suffices to show that `\Gamma', y:A \vdash t' : t'`.
By the IH (setting $$\Gamma'' = \Gamma', y:A$$), it suffices to show By the IH (setting `\Gamma'' = \Gamma', y:A`), it suffices to show
that $$\Gamma, y:A$$ and $$\Gamma', y:A$$ agree on all the variables that `\Gamma, y:A` and `\Gamma', y:A` agree on all the variables
that appear free in $$t'$$. that appear free in `t'`.
Any variable occurring free in $$t'$$ must be either $$y$$ or Any variable occurring free in `t'` must be either `y` or
some other variable. $$\Gamma, y:A$$ and $$\Gamma', y:A$$ some other variable. `\Gamma, y:A` and `\Gamma', y:A`
clearly agree on $$y$$. Otherwise, note that any variable other clearly agree on `y`. Otherwise, note that any variable other
than $$y$$ that occurs free in $$t'$$ also occurs free in than `y` that occurs free in `t'` also occurs free in
$$t = \lambda y:A. t'$$, and by assumption $$\Gamma$$ and `t = \lambda y:A. t'`, and by assumption `\Gamma` and
$$\Gamma'$$ agree on all such variables; hence so do $$\Gamma, y:A$$ and `\Gamma'` agree on all such variables; hence so do `\Gamma, y:A` and
$$\Gamma', y:A$$. `\Gamma', y:A`.
- If the last rule was `app`, then $$t = t_1\;t_2$$, with - If the last rule was `app`, then `t = t_1\;t_2`, with
$$\Gamma \vdash t_1:A\to T$$ and $$\Gamma \vdash t_2:A$$. `\Gamma \vdash t_1:A\to T` and `\Gamma \vdash t_2:A`.
One induction hypothesis states that for all contexts $$\Gamma'$$, One induction hypothesis states that for all contexts `\Gamma'`,
if $$\Gamma'$$ agrees with $$\Gamma$$ on the free variables in $$t_1$$, if `\Gamma'` agrees with `\Gamma` on the free variables in `t_1`,
then $$t_1$$ has type $$A\to T$$ under $$\Gamma'$$; there is a similar IH then `t_1` has type `A\to T` under `\Gamma'`; there is a similar IH
for $$t_2$$. We must show that $$t_1\;t_2$$ also has type $$T$$ under for `t_2`. We must show that `t_1\;t_2` also has type `T` under
$$\Gamma'$$, given the assumption that $$\Gamma'$$ agrees with `\Gamma'`, given the assumption that `\Gamma'` agrees with
$$\Gamma$$ on all the free variables in $$t_1\;t_2$$. By `app`, it `\Gamma` on all the free variables in `t_1\;t_2`. By `app`, it
suffices to show that $$t_1$$ and $$t_2$$ each have the same type suffices to show that `t_1` and `t_2` each have the same type
under $$\Gamma'$$ as under $$\Gamma$$. But all free variables in under `\Gamma'` as under `\Gamma`. But all free variables in
$$t_1$$ are also free in $$t_1\;t_2$$, and similarly for $$t_2$$; `t_1` are also free in `t_1\;t_2`, and similarly for `t_2`;
hence the desired result follows from the induction hypotheses. hence the desired result follows from the induction hypotheses.
\begin{code} \begin{code}
@ -383,16 +383,16 @@ preserves types---namely, the observation that _substitution_
preserves types. preserves types.
Formally, the so-called _Substitution Lemma_ says this: Suppose we Formally, the so-called _Substitution Lemma_ says this: Suppose we
have a term $$N$$ with a free variable $$x$$, and suppose we've been 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 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 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 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 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$$ able to substitute `V` for each of the occurrences of `x` in `N`
and obtain a new term that still has type $$B$$. and obtain a new term that still has type `B`.
_Lemma_: If $$Γ , x ↦ A ⊢ N ∈ B$$ and $$∅ ⊢ V ∈ A$$, then _Lemma_: If `Γ , x ↦ A ⊢ N ∈ B` and `∅ ⊢ V ∈ A`, then
$$Γ ⊢ (N [ x := V ]) ∈ B$$. `Γ ⊢ (N [ x := V ]) ∈ B`.
\begin{code} \begin{code}
preservation-[:=] : ∀ {Γ x A N B V} preservation-[:=] : ∀ {Γ x A N B V}
@ -402,63 +402,63 @@ preservation-[:=] : ∀ {Γ x A N B V}
\end{code} \end{code}
One technical subtlety in the statement of the lemma is that One technical subtlety in the statement of the lemma is that
we assign $$V$$ the type $$A$$ in the _empty_ context---in other we assign `V` the type `A` in the _empty_ context---in other
words, we assume $$V$$ is closed. This assumption considerably words, we assume `V` is closed. This assumption considerably
simplifies the $$λᵀ$$ case of the proof (compared to assuming simplifies the `λᵀ` case of the proof (compared to assuming
$$Γ ⊢ V ∈ A$$, which would be the other reasonable assumption `Γ ⊢ V ∈ A`, which would be the other reasonable assumption
at this point) because the context invariance lemma then tells us at this point) because the context invariance lemma then tells us
that $$V$$ has type $$A$$ in any context at all---we don't have to 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 worry about free variables in `V` clashing with the variable being
introduced into the context by $$λᵀ$$. introduced into the context by `λᵀ`.
The substitution lemma can be viewed as a kind of "commutation" The substitution lemma can be viewed as a kind of "commutation"
property. Intuitively, it says that substitution and typing can property. Intuitively, it says that substitution and typing can
be done in either order: we can either assign types to the terms be done in either order: we can either assign types to the terms
$$N$$ 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 them using substitution, or we can substitute first and then
assign a type to $$N [ x := V ]$$---the result is the same either assign a type to `N [ x := V ]`---the result is the same either
way. way.
_Proof_: We show, by induction on $$N$$, that for all $$A$$ and _Proof_: We show, by induction on `N`, that for all `A` and
$$Γ$$, if $$Γ , x ↦ A \vdash N ∈ B$$ and $$∅ ⊢ V ∈ A$$, then `Γ`, if `Γ , x ↦ A \vdash N ∈ B` and `∅ ⊢ V ∈ A`, then
$$Γ \vdash N [ x := V ] ∈ B$$. `Γ \vdash N [ x := V ] ∈ B`.
- If $$N$$ is a variable there are two cases to consider, - If `N` is a variable there are two cases to consider,
depending on whether $$N$$ is $$x$$ or some other variable. depending on whether `N` is `x` or some other variable.
- If $$N = varᵀ x$$, then from the fact that $$Γ , x ↦ A ⊢ N ∈ B$$ - 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] = 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 `Γ`, given the assumption that
$$V$$ has type $$A$$ under the empty context. This `V` has type `A` under the empty context. This
follows from context invariance: if a closed term has type follows from context invariance: if a closed term has type
$$A$$ in the empty context, it has that type in any context. `A` in the empty context, it has that type in any context.
- If $$N$$ is some variable $$x$$ different from $$x$$, then - If `N` is some variable `x` different from `x`, then
we need only note that $$x$$ has the same type under $$Γ , x ↦ A$$ we need only note that `x` has the same type under `Γ , x ↦ A`
as under $$Γ$$. as under `Γ`.
- If $$N$$ is an abstraction $$λᵀ x ∈ A ⇒ N$$, then the IH tells us, - If `N` is an abstraction `λᵀ x ∈ A ⇒ N`, then the IH tells us,
for all $$Γ′$$́ and $$B$$, that if $$Γ′ , x ↦ A ⊢ N ∈ B$$ for all `Γ′`́ and `B`, that if `Γ′ , x ↦ A ⊢ N ∈ B`
and $$∅ ⊢ V ∈ A$$, then $$Γ′ ⊢ N [ x := V ] ∈ B$$. and `∅ ⊢ V ∈ A`, then `Γ′ ⊢ N [ x := V ] ∈ B`.
The substitution in the conclusion behaves differently The substitution in the conclusion behaves differently
depending on whether $$x$$ and $$x$$ are the same variable. depending on whether `x` and `x` are the same variable.
First, suppose $$x ≡ x$$. Then, by the definition of First, suppose `x ≡ x`. Then, by the definition of
substitution, $$N [ x := V] = N$$, so we just need to show $$Γ ⊢ N ∈ B$$. substitution, `N [ x := V] = N`, so we just need to show `Γ ⊢ N ∈ B`.
But we know $$Γ , x ↦ A ⊢ N ∈ B$$ and, since $$x ≡ x$$ But we know `Γ , x ↦ A ⊢ N ∈ B` and, since `x ≡ x`
does not appear free in $$λᵀ x ∈ A ⇒ N$$, the context invariance does not appear free in `λᵀ x ∈ A ⇒ N`, the context invariance
lemma yields $$Γ ⊢ N ∈ B$$. lemma yields `Γ ⊢ N ∈ B`.
Second, suppose $$x ≢ x$$. We know $$Γ , x ↦ A , x ↦ A ⊢ N ∈ B$$ Second, suppose `x ≢ x`. We know `Γ , x ↦ A , x ↦ A ⊢ N ∈ B`
by inversion of the typing relation, from which by inversion of the typing relation, from which
$$Γ , x ↦ A , x ↦ A ⊢ N ∈ B$$ follows by update permute, `Γ , x ↦ A , x ↦ A ⊢ N ∈ B` follows by update permute,
so the IH applies, giving us $$Γ , x ↦ A ⊢ N [ x := V ] ∈ B$$ so the IH applies, giving us `Γ , x ↦ A ⊢ N [ x := V ] ∈ B`
By $$⇒-I$$, we have $$Γ ⊢ λᵀ x ∈ A ⇒ (N [ x := V ]) ∈ A ⇒ B$$ By `⇒-I`, we have `Γ ⊢ λᵀ x ∈ A ⇒ (N [ x := V ]) ∈ A ⇒ B`
and the definition of substitution (noting $$x ≢ x$$) gives and the definition of substitution (noting `x ≢ x`) gives
$$Γ ⊢ (λᵀ x ∈ A ⇒ N) [ x := V ] ∈ A ⇒ B$$ as required. `Γ ⊢ (λᵀ x ∈ A ⇒ N) [ x := V ] ∈ A ⇒ B` as required.
- If $$N$$ is an application $$L ·ᵀ M$$, the result follows - If `N` is an application `L ·ᵀ M`, the result follows
straightforwardly from the definition of substitution and the straightforwardly from the definition of substitution and the
induction hypotheses. induction hypotheses.
@ -483,12 +483,7 @@ just-injective refl = refl
preservation-[:=] {_} {x} (Ax {_} {x} [Γ,x↦A]x≡B) ⊢V with x ≟ x 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 ...| yes x≡x rewrite just-injective [Γ,x↦A]x≡B = weaken-closed ⊢V
...| no x≢x = Ax [Γ,x↦A]x≡B ...| no x≢x = Ax [Γ,x↦A]x≡B
{- preservation-[:=] {Γ} {x} {A} {λᵀ x ∈ A ⇒ N} {.A ⇒ B} {V} (⇒-I ⊢N) ⊢V with x ≟ x
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
...| yes x≡x rewrite x≡x = weaken Γ′~Γ (⇒-I ⊢N) ...| yes x≡x rewrite x≡x = weaken Γ′~Γ (⇒-I ⊢N)
where where
Γ′~Γ : ∀ {y} → y FreeIn (λᵀ x ∈ A ⇒ N) → (Γ , x ↦ A) y ≡ Γ y Γ′~Γ : ∀ {y} → y FreeIn (λᵀ x ∈ A ⇒ N) → (Γ , x ↦ A) y ≡ Γ y
@ -497,79 +492,58 @@ preservation-[:=] {Γ} {x} {A} {λᵀ x ∈ A ⇒ N} {.A ⇒ B} {
...| no _ = refl ...| no _ = refl
...| no x≢x = ⇒-I ⊢NV ...| no x≢x = ⇒-I ⊢NV
where where
xx⊢N : (Γ , x ↦ A , x ↦ A) ⊢ N ∈ B xx⊢N : Γ , x ↦ A , x ↦ A ⊢ N ∈ B
xx⊢N rewrite update-permute Γ x A x A x≢x = {!⊢N!} xx⊢N rewrite update-permute Γ x A x A x≢x = ⊢N
⊢NV : (Γ , x ↦ A) ⊢ N [ x := V ] ∈ B ⊢NV : (Γ , x ↦ A) ⊢ N [ x := V ] ∈ B
⊢NV = preservation-[:=] xx⊢N ⊢V ⊢NV = preservation-[:=] xx⊢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-[:=] (⇒-E ⊢L ⊢M) ⊢V = ⇒-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V)
preservation-[:=] 𝔹-I₁ ⊢V = 𝔹-I₁ preservation-[:=] 𝔹-I₁ ⊢V = 𝔹-I₁
preservation-[:=] 𝔹-I₂ ⊢V = 𝔹-I₂ preservation-[:=] 𝔹-I₂ ⊢V = 𝔹-I₂
preservation-[:=] (𝔹-E ⊢L ⊢M ⊢N) ⊢V = preservation-[:=] (𝔹-E ⊢L ⊢M ⊢N) ⊢V =
𝔹-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) (preservation-[:=] ⊢N ⊢V) 𝔹-E (preservation-[:=] ⊢L ⊢V) (preservation-[:=] ⊢M ⊢V) (preservation-[:=] ⊢N ⊢V)
{-
[:=]-preserves-⊢ {Γ} {x} vA (var y y∈Γ) with x ≟ y
... | yes x=y = {!!}
... | no x≠y = {!!}
[:=]-preserves-⊢ vA (abs tB) = {!!}
[:=]-preserves-⊢ vA (app t₁A⇒B t₂A) =
app ([:=]-preserves-⊢ vA t₁A⇒B) ([:=]-preserves-⊢ vA t₂A)
[:=]-preserves-⊢ vA true = true
[:=]-preserves-⊢ vA false = false
[:=]-preserves-⊢ vA (if t₁bool then t₂B else t₃B) =
if [:=]-preserves-⊢ vA t₁bool
then [:=]-preserves-⊢ vA t₂B
else [:=]-preserves-⊢ vA t₃B
-}
\end{code} \end{code}
### Main Theorem ### Main Theorem
We now have the tools we need to prove preservation: if a closed We now have the tools we need to prove preservation: if a closed
term $$M$$ has type $$A$$ and takes a step to $$N$$, then $$N$$ 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 is also a closed term with type `A`. In other words, small-step
reduction preserves types. reduction preserves types.
\begin{code} \begin{code}
preservation : ∀ {M N A} → ∅ ⊢ M ∈ A → M ⟹ N → ∅ ⊢ N ∈ A preservation : ∀ {M N A} → ∅ ⊢ M ∈ A → M ⟹ N → ∅ ⊢ N ∈ A
\end{code} \end{code}
_Proof_: By induction on the derivation of $$\vdash t : T$$. _Proof_: By induction on the derivation of `\vdash t : T`.
- We can immediately rule out $$var$$, $$abs$$, $$T_True$$, and - We can immediately rule out `var`, `abs`, `T_True`, and
$$T_False$$ as the final rules in the derivation, since in each of `T_False` as the final rules in the derivation, since in each of
these cases $$t$$ cannot take a step. these cases `t` cannot take a step.
- If the last rule in the derivation was $$app$$, then $$t = t_1 - If the last rule in the derivation was `app`, then `t = t_1
t_2$$. There are three cases to consider, one for each rule that t_2`. There are three cases to consider, one for each rule that
could have been used to show that $$t_1 t_2$$ takes a step to $$t'$$. could have been used to show that `t_1 t_2` takes a step to `t'`.
- If $$t_1 t_2$$ takes a step by $$Sapp1$$, with $$t_1$$ stepping to - If `t_1 t_2` takes a step by `Sapp1`, with `t_1` stepping to
$$t_1'$$, then by the IH $$t_1'$$ has the same type as $$t_1$$, and `t_1'`, then by the IH `t_1'` has the same type as `t_1`, and
hence $$t_1' t_2$$ has the same type as $$t_1 t_2$$. hence `t_1' t_2` has the same type as `t_1 t_2`.
- The $$Sapp2$$ case is similar. - The `Sapp2` case is similar.
- If $$t_1 t_2$$ takes a step by $$Sred$$, then $$t_1 = - If `t_1 t_2` takes a step by `Sred`, then `t_1 =
\lambda x:t_{11}.t_{12}$$ and $$t_1 t_2$$ steps to $$$$x:=t_2$$t_{12}$$; the \lambda x:t_{11}.t_{12}` and `t_1 t_2` steps to ``x:=t_2`t_{12}`; the
desired result now follows from the fact that substitution desired result now follows from the fact that substitution
preserves types. preserves types.
- If the last rule in the derivation was $$if$$, then $$t = if t_1 - If the last rule in the derivation was `if`, then `t = if t_1
then t_2 else t_3$$, and there are again three cases depending on then t_2 else t_3`, and there are again three cases depending on
how $$t$$ steps. how `t` steps.
- If $$t$$ steps to $$t_2$$ or $$t_3$$, the result is immediate, since - If `t` steps to `t_2` or `t_3`, the result is immediate, since
$$t_2$$ and $$t_3$$ have the same type as $$t$$. `t_2` and `t_3` have the same type as `t`.
- Otherwise, $$t$$ steps by $$Sif$$, and the desired conclusion - Otherwise, `t` steps by `Sif`, and the desired conclusion
follows directly from the induction hypothesis. follows directly from the induction hypothesis.
\begin{code} \begin{code}
@ -597,11 +571,11 @@ Proof with eauto.
intros t t' T HT. generalize dependent t'. intros t t' T HT. generalize dependent t'.
induction HT; induction HT;
intros t' HE; subst Gamma; subst; intros t' HE; subst Gamma; subst;
try solve $$inversion HE; subst; auto$$. try solve `inversion HE; subst; auto`.
- (* app - (* app
inversion HE; subst... inversion HE; subst...
(* Most of the cases are immediate by induction, (* Most of the cases are immediate by induction,
and $$eauto$$ takes care of them and `eauto` takes care of them
+ (* Sred + (* Sred
apply substitution_preserves_typing with t_{11}... apply substitution_preserves_typing with t_{11}...
inversion HT_1... inversion HT_1...
@ -611,7 +585,7 @@ Qed.
An exercise in the [Stlc]({{ "Stlc" | relative_url }}) chapter asked about the An exercise in the [Stlc]({{ "Stlc" | relative_url }}) chapter asked about the
subject expansion property for the simple language of arithmetic and boolean subject expansion property for the simple language of arithmetic and boolean
expressions. Does this property hold for STLC? That is, is it always the case 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 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. so, prove it. If not, give a counter-example not involving conditionals.
@ -630,7 +604,7 @@ Corollary soundness : forall t t' T,
~(stuck t'). ~(stuck t').
Proof. Proof.
intros t t' T Hhas_type Hmulti. unfold stuck. intros t t' T Hhas_type Hmulti. unfold stuck.
intros $$Hnf Hnot_val$$. unfold normal_form in Hnf. intros `Hnf Hnot_val`. unfold normal_form in Hnf.
induction Hmulti. induction Hmulti.
@ -647,10 +621,10 @@ Formalize this statement and prove it.
#### Exercise: 1 star (progress_preservation_statement) #### Exercise: 1 star (progress_preservation_statement)
Without peeking at their statements above, write down the progress Without peeking at their statements above, write down the progress
and preservation theorems for the simply typed lambda-calculus. and preservation theorems for the simply typed lambda-calculus.
$$$$ ``
#### Exercise: 2 stars (stlc_variation1) #### Exercise: 2 stars (stlc_variation1)
Suppose we add a new term $$zap$$ with the following reduction rule Suppose we add a new term `zap` with the following reduction rule
--------- (ST_Zap) --------- (ST_Zap)
t ==> zap t ==> zap
@ -665,7 +639,7 @@ the presence of these rules? For each property, write either
"remains true" or "becomes false." If a property becomes "remains true" or "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -673,7 +647,7 @@ false, give a counterexample.
#### Exercise: 2 stars (stlc_variation2) #### Exercise: 2 stars (stlc_variation2)
Suppose instead that we add a new term $$foo$$ with the following Suppose instead that we add a new term `foo` with the following
reduction rules: reduction rules:
----------------- (ST_Foo1) ----------------- (ST_Foo1)
@ -687,20 +661,20 @@ the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
- Preservation - Preservation
#### Exercise: 2 stars (stlc_variation3) #### Exercise: 2 stars (stlc_variation3)
Suppose instead that we remove the rule $$Sapp1$$ from the $$step$$ Suppose instead that we remove the rule `Sapp1` from the `step`
relation. Which of the following properties of the STLC remain relation. Which of the following properties of the STLC remain
true in the presence of this rule? For each one, write either true in the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -718,7 +692,7 @@ the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -740,7 +714,7 @@ the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -762,7 +736,7 @@ the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -782,7 +756,7 @@ the presence of this rule? For each one, write either
"remains true" or else "becomes false." If a property becomes "remains true" or else "becomes false." If a property becomes
false, give a counterexample. false, give a counterexample.
- Determinism of $$step$$ - Determinism of `step`
- Progress - Progress
@ -824,10 +798,10 @@ with arithmetic. Specifically:
the definition of values through the Type Soundness theorem), and the definition of values through the Type Soundness theorem), and
paste it into the file at this point. paste it into the file at this point.
- Extend the definitions of the $$subst$$ operation and the $$step$$ - Extend the definitions of the `subst` operation and the `step`
relation to include appropriate clauses for the arithmetic operators. relation to include appropriate clauses for the arithmetic operators.
- Extend the proofs of all the properties (up to $$soundness$$) of - Extend the proofs of all the properties (up to `soundness`) of
the original STLC to deal with the new syntactic forms. Make the original STLC to deal with the new syntactic forms. Make
sure Agda accepts the whole file. sure Agda accepts the whole file.