failed attempt to use std lib reasoning

This commit is contained in:
Jeremy Siek 2019-08-20 15:31:52 -04:00
parent 3f9f9c3c00
commit b88cd2478b
2 changed files with 99 additions and 75 deletions

View file

@ -94,10 +94,12 @@ sub- {v = v₁ ↦ v₂ ⊔ v₁ ↦ v₃} {v₁ ↦ (v₂ ⊔ v₃)} ⟨ N2
sub- d (⊑-trans x₁ x₂) = sub- (sub- d x₂) x₁
```
<!--
[PLW:
If denotations were strengthened to be downward closed,
we could rewrite the signature replacing ( N) by d : Denotation (Γ , ★)]
[JGS: I'll look into this.]
-->
With this subsumption property in hand, we can prove the forward
direction of the semantic equation for lambda. The proof is by

View file

@ -54,16 +54,19 @@ down a denotational semantics of the lambda calculus.
## Imports
<!-- JGS: for equational reasoning
open import Relation.Binary using (Setoid)
-->
```
open import Relation.Binary.PropositionalEquality
using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app)
open import Agda.Primitive using (lzero; lsuc)
open import Data.Empty using (⊥-elim)
open import Data.Product using (_×_; Σ; Σ-syntax; ∃; ∃-syntax; proj₁; proj₂)
renaming (_,_ to ⟨_,_⟩)
open import Data.Sum
open import Agda.Primitive using (lzero)
open import Relation.Binary.PropositionalEquality
using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app)
open import Relation.Nullary using (¬_)
open import Relation.Nullary.Negation using (contradiction)
open import Data.Empty using (⊥-elim)
open import Function using (_∘_)
open import plfa.part2.Untyped
using (Context; ★; _∋_; ∅; _,_; Z; S_; _⊢_; `_; _·_; ƛ_;
@ -187,7 +190,9 @@ property.
(⊑-fun (⊑-conj-R2 ⊑-refl) ⊑-refl))
```
<!-- above might read more nicely if we introduce inequational reasoning -->
<!--
[PLW: above might read more nicely if we introduce inequational reasoning.]
-->
If the join `u ⊔ v` is less than another value `w`,
then both `u` and `v` are less than `w`.
@ -246,8 +251,7 @@ last γ = γ Z
init-last : ∀ {Γ} → (γ : Env (Γ , ★)) → γ ≡ (init γ `, last γ)
init-last {Γ} γ = extensionality lemma
where
lemma : ∀ (x : Γ , ★ ∋ ★) → γ x ≡ (init γ `, last γ) x
where lemma : ∀ (x : Γ , ★ ∋ ★) → γ x ≡ (init γ `, last γ) x
lemma Z = refl
lemma (S x) = refl
```
@ -576,7 +580,25 @@ equal, that is, ` M ≃ N`.
The following submodule introduces equational reasoning for the `≃`
relation.
<!--
JGS: I couldn't get this to work. The definitions here were accepted
by Agda, but then the uses in the Compositional chapter got rejected.
denotation-setoid : Context → Setoid (lsuc lzero) lzero
denotation-setoid Γ = record
{ Carrier = Denotation Γ
; _≈_ = _≃_
; isEquivalence = record { refl = ≃-refl ; sym = ≃-sym ; trans = ≃-trans } }
-->
<!--
The following went inside the module ≃-Reasoning:
open import Relation.Binary.Reasoning.Setoid (denotation-setoid Γ)
renaming (begin_ to start_; _≈⟨_⟩_ to _≃⟨_⟩_; _∎ to _☐) public
-->
```
module ≃-Reasoning {Γ : Context} where
infix 1 start_
@ -589,12 +611,6 @@ module ≃-Reasoning {Γ : Context} where
→ x ≃ y
start x≃y = x≃y
_≃⟨⟩_ : ∀ (x : Denotation Γ) {y : Denotation Γ}
→ x ≃ y
-----
→ x ≃ y
x ≃⟨⟩ x≃y = x≃y
_≃⟨_⟩_ : ∀ (x : Denotation Γ) {y z : Denotation Γ}
→ x ≃ y
→ y ≃ z
@ -602,6 +618,12 @@ module ≃-Reasoning {Γ : Context} where
→ x ≃ z
(x ≃⟨ x≃y ⟩ y≃z) = ≃-trans x≃y y≃z
_≃⟨⟩_ : ∀ (x : Denotation Γ) {y : Denotation Γ}
→ x ≃ y
-----
→ x ≃ y
x ≃⟨⟩ x≃y = x≃y
_☐ : ∀ (x : Denotation Γ)
-----
→ x ≃ x