updated LambdaProp to PCF
This commit is contained in:
parent
ebc9832f29
commit
f9e21d7ecf
1 changed files with 32 additions and 28 deletions
|
@ -312,7 +312,6 @@ subst {x = y} (ℕ-E {x = x} ⊢L ⊢M ⊢N) ⊢V with x ≟ y
|
||||||
subst {x = y} (Fix {x = x} ⊢M) ⊢V with x ≟ y
|
subst {x = y} (Fix {x = x} ⊢M) ⊢V with x ≟ y
|
||||||
... | yes refl = Fix (rename-≡ ⊢M)
|
... | yes refl = Fix (rename-≡ ⊢M)
|
||||||
... | no x≢y = Fix (subst (rename-≢ x≢y ⊢M) ⊢V)
|
... | no x≢y = Fix (subst (rename-≢ x≢y ⊢M) ⊢V)
|
||||||
{-
|
|
||||||
\end{code}
|
\end{code}
|
||||||
|
|
||||||
|
|
||||||
|
@ -324,20 +323,22 @@ 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
|
preserve : ∀ {M N A}
|
||||||
preservation (Ax ())
|
→ ∅ ⊢ M ⦂ A
|
||||||
preservation (⇒-I ⊢N) ()
|
→ M ⟹ N
|
||||||
preservation (⇒-E ⊢L ⊢M) (ξ·₁ L⟹L′) with preservation ⊢L L⟹L′
|
----------
|
||||||
... | ⊢L′ = ⇒-E ⊢L′ ⊢M
|
→ ∅ ⊢ N ⦂ A
|
||||||
preservation (⇒-E ⊢L ⊢M) (ξ·₂ valueL M⟹M′) with preservation ⊢M M⟹M′
|
preserve (Ax ())
|
||||||
... | ⊢M′ = ⇒-E ⊢L ⊢M′
|
preserve (⇒-I ⊢N) ()
|
||||||
preservation (⇒-E (⇒-I ⊢N) ⊢V) (βλ· valueV) = subst ⊢N ⊢V
|
preserve (⇒-E ⊢L ⊢M) (ξ-·₁ L⟹L′) = ⇒-E (preserve ⊢L L⟹L′) ⊢M
|
||||||
preservation `ℕ-I₁ ()
|
preserve (⇒-E ⊢L ⊢M) (ξ-·₂ VL M⟹M′) = ⇒-E ⊢L (preserve ⊢M M⟹M′)
|
||||||
preservation `ℕ-I₂ ()
|
preserve (⇒-E (⇒-I ⊢N) ⊢V) (β-ƛ· VV) = subst ⊢N ⊢V
|
||||||
preservation (`ℕ-E ⊢L ⊢M ⊢N) (ξif L⟹L′) with preservation ⊢L L⟹L′
|
preserve ℕ-I₁ ()
|
||||||
... | ⊢L′ = `ℕ-E ⊢L′ ⊢M ⊢N
|
preserve (ℕ-I₂ ⊢M) (ξ-suc M⟹M′) = ℕ-I₂ (preserve ⊢M M⟹M′)
|
||||||
preservation (`ℕ-E `ℕ-I₁ ⊢M ⊢N) βif-true = ⊢M
|
preserve (ℕ-E ⊢L ⊢M ⊢N) (ξ-case L⟹L′) = ℕ-E (preserve ⊢L L⟹L′) ⊢M ⊢N
|
||||||
preservation (`ℕ-E `ℕ-I₂ ⊢M ⊢N) βif-false = ⊢N
|
preserve (ℕ-E ℕ-I₁ ⊢M ⊢N) β-case-zero = ⊢M
|
||||||
|
preserve (ℕ-E (ℕ-I₂ ⊢V) ⊢M ⊢N) (β-case-suc VV) = subst ⊢N ⊢V
|
||||||
|
preserve (Fix ⊢M) (β-μ) = subst ⊢M (Fix ⊢M)
|
||||||
\end{code}
|
\end{code}
|
||||||
|
|
||||||
|
|
||||||
|
@ -357,7 +358,7 @@ The preservation property is sometimes called _subject reduction_.
|
||||||
Its opposite is _subject expansion_, which holds if
|
Its opposite is _subject expansion_, which holds if
|
||||||
`M ==> N` and `∅ ⊢ N ⦂ A`, then `∅ ⊢ M ⦂ A`.
|
`M ==> N` and `∅ ⊢ N ⦂ A`, then `∅ ⊢ M ⦂ A`.
|
||||||
Find two counter-examples to subject expansion, one
|
Find two counter-examples to subject expansion, one
|
||||||
with conditionals and one not involving conditionals.
|
with case expressions and one not involving case expressions.
|
||||||
|
|
||||||
## Type Soundness
|
## Type Soundness
|
||||||
|
|
||||||
|
@ -374,18 +375,24 @@ Stuck : Term → Set
|
||||||
Stuck M = Normal M × ¬ Value M
|
Stuck M = Normal M × ¬ Value M
|
||||||
|
|
||||||
postulate
|
postulate
|
||||||
Soundness : ∀ {M N A} → ∅ ⊢ M ⦂ A → M ⟹* N → ¬ (Stuck N)
|
Soundness : ∀ {M N A}
|
||||||
|
→ ∅ ⊢ M ⦂ A
|
||||||
|
→ M ⟹* N
|
||||||
|
-----------
|
||||||
|
→ ¬ (Stuck N)
|
||||||
\end{code}
|
\end{code}
|
||||||
|
|
||||||
<div class="hidden">
|
<div class="hidden">
|
||||||
\begin{code}
|
\begin{code}
|
||||||
Soundness′ : ∀ {M N A} → ∅ ⊢ M ⦂ A → M ⟹* N → ¬ (Stuck N)
|
Soundness′ : ∀ {M N A}
|
||||||
Soundness′ ⊢M (M ∎) ⟨ ¬M⟹N , ¬ValueM ⟩ with progress ⊢M
|
→ ∅ ⊢ M ⦂ A
|
||||||
... | steps M⟹N = ¬M⟹N M⟹N
|
→ M ⟹* N
|
||||||
... | done ValueM = ¬ValueM ValueM
|
-----------
|
||||||
Soundness′ ⊢L (L ⟹⟨ L⟹M ⟩ M⟹*N) = Soundness′ ⊢M M⟹*N
|
→ ¬ (Stuck N)
|
||||||
where
|
Soundness′ ⊢M (M ∎) ⟨ ¬M⟹N , ¬VM ⟩ with progress ⊢M
|
||||||
⊢M = preservation ⊢L L⟹M
|
... | steps M⟹N = ¬M⟹N M⟹N
|
||||||
|
... | done VM = ¬VM VM
|
||||||
|
Soundness′ ⊢L (L ⟹⟨ L⟹M ⟩ M⟹*N) = Soundness′ (preserve ⊢L L⟹M) M⟹*N
|
||||||
\end{code}
|
\end{code}
|
||||||
</div>
|
</div>
|
||||||
|
|
||||||
|
@ -527,7 +534,7 @@ Suppose we add the following new rule to the typing relation
|
||||||
of the STLC:
|
of the STLC:
|
||||||
|
|
||||||
-------------------- (T_FunnyAbs)
|
-------------------- (T_FunnyAbs)
|
||||||
∅ ⊢ λ[ x ⦂ `ℕ ] N ⦂ `ℕ
|
∅ ⊢ ƛ x ⇒ N ⦂ `ℕ
|
||||||
|
|
||||||
Which of the following properties of the STLC remain true in
|
Which of the following properties of the STLC remain true in
|
||||||
the presence of this rule? For each one, write either
|
the presence of this rule? For each one, write either
|
||||||
|
@ -541,6 +548,3 @@ false, give a counterexample.
|
||||||
- Preservation
|
- Preservation
|
||||||
|
|
||||||
|
|
||||||
\begin{code}
|
|
||||||
-}
|
|
||||||
\end{code}
|
|
||||||
|
|
Loading…
Reference in a new issue