moving denotational back to plfa, updates to Confluence

This commit is contained in:
Jeremy Siek 2019-05-25 16:11:01 -04:00
parent fcfcd2a70c
commit 18cec9a0da
9 changed files with 30 additions and 31 deletions

View file

@ -7,7 +7,7 @@ next : /ContextualEquivalence/
--- ---
\begin{code} \begin{code}
module denotational.Adequacy where module plfa.Adequacy where
\end{code} \end{code}
## Imports ## Imports
@ -16,18 +16,18 @@ module denotational.Adequacy where
open import plfa.Untyped open import plfa.Untyped
using (Context; _⊢_; ★; _∋_; ∅; _,_; Z; S_; `_; ƛ_; _·_; using (Context; _⊢_; ★; _∋_; ∅; _,_; Z; S_; `_; ƛ_; _·_;
rename; subst; ext; exts; _[_]; subst-zero) rename; subst; ext; exts; _[_]; subst-zero)
open import denotational.LambdaReduction open import plfa.LambdaReduction
using (_—↠_; _—→⟨_⟩_; _[]; _—→_; ξ₁; ξ₂; β; ζ) using (_—↠_; _—→⟨_⟩_; _[]; _—→_; ξ₁; ξ₂; β; ζ)
open import denotational.CallByName open import plfa.CallByName
using (Clos; clos; ClosEnv; ∅'; _,'_; _⊢_⇓_; ⇓-var; ⇓-lam; ⇓-app; ⇓-determ; using (Clos; clos; ClosEnv; ∅'; _,'_; _⊢_⇓_; ⇓-var; ⇓-lam; ⇓-app; ⇓-determ;
cbn→reduce) cbn→reduce)
open import denotational.Denotational open import plfa.Denotational
using (Value; Env; `∅; _`,_; _↦_; _⊑_; _⊢_↓_; ⊥; Funs∈; _⊔_; ∈→⊑; using (Value; Env; `∅; _`,_; _↦_; _⊑_; _⊢_↓_; ⊥; Funs∈; _⊔_; ∈→⊑;
var; ↦-elim; ↦-intro; ⊔-intro; ⊥-intro; sub; ; _≃_; _iff_; var; ↦-elim; ↦-intro; ⊔-intro; ⊥-intro; sub; ; _≃_; _iff_;
Trans⊑; ConjR1⊑; ConjR2⊑; ConjL⊑; Refl⊑; Fun⊑; Bot⊑; Dist⊑; Trans⊑; ConjR1⊑; ConjR2⊑; ConjL⊑; Refl⊑; Fun⊑; Bot⊑; Dist⊑;
sub-inv-fun) sub-inv-fun)
open import denotational.Soundness using (soundness) open import plfa.Soundness using (soundness)
open import denotational.Substitution using (ids; sub-id) open import plfa.Substitution using (ids; sub-id)
import Relation.Binary.PropositionalEquality as Eq import Relation.Binary.PropositionalEquality as Eq
open Eq using (_≡_; _≢_; refl; trans; sym; cong; cong₂; cong-app) open Eq using (_≡_; _≢_; refl; trans; sym; cong; cong₂; cong-app)

View file

@ -7,7 +7,7 @@ next : /Denotational/
--- ---
\begin{code} \begin{code}
module denotational.CallByName where module plfa.CallByName where
\end{code} \end{code}
## Imports ## Imports
@ -16,9 +16,9 @@ module denotational.CallByName where
open import plfa.Untyped open import plfa.Untyped
using (Context; _⊢_; _∋_; ★; ∅; _,_; Z; S_; `_; ƛ_; _·_; subst; subst-zero; using (Context; _⊢_; _∋_; ★; ∅; _,_; Z; S_; `_; ƛ_; _·_; subst; subst-zero;
exts; rename) exts; rename)
open import denotational.LambdaReduction open import plfa.LambdaReduction
using (β; ξ₁; ξ₂; ζ; _—→_; _—↠_; _—→⟨_⟩_; _[]; —↠-trans; appL-cong) using (β; ξ₁; ξ₂; ζ; _—→_; _—↠_; _—→⟨_⟩_; _[]; —↠-trans; appL-cong)
open import denotational.Substitution open import plfa.Substitution
using (Subst; ⟪_⟫; _•_; _⨟_; ids; sub-id; sub-sub; subst-zero-exts-cons) using (Subst; ⟪_⟫; _•_; _⨟_; ids; sub-id; sub-sub; subst-zero-exts-cons)
import Relation.Binary.PropositionalEquality as Eq import Relation.Binary.PropositionalEquality as Eq
open Eq using (_≡_; refl; trans; sym) open Eq using (_≡_; refl; trans; sym)

View file

@ -7,7 +7,7 @@ next : /Soundness/
--- ---
\begin{code} \begin{code}
module denotational.Compositional where module plfa.Compositional where
\end{code} \end{code}
## Imports ## Imports
@ -15,12 +15,12 @@ module denotational.Compositional where
\begin{code} \begin{code}
open import plfa.Untyped open import plfa.Untyped
using (Context; _,_; ★; _∋_; _⊢_; `_; ƛ_; _·_) using (Context; _,_; ★; _∋_; _⊢_; `_; ƛ_; _·_)
open import denotational.Denotational open import plfa.Denotational
using (Value; _↦_; _`,_; _⊔_; ⊥; _⊑_; _⊢_↓_; nth; using (Value; _↦_; _`,_; _⊔_; ⊥; _⊑_; _⊢_↓_; nth;
Bot⊑; Fun⊑; ConjL⊑; ConjR1⊑; ConjR2⊑; Dist⊑; Refl⊑; Trans⊑; Dist⊔↦⊔; Bot⊑; Fun⊑; ConjL⊑; ConjR1⊑; ConjR2⊑; Dist⊑; Refl⊑; Trans⊑; Dist⊔↦⊔;
var; ↦-intro; ↦-elim; ⊔-intro; ⊥-intro; sub; var; ↦-intro; ↦-elim; ⊔-intro; ⊥-intro; sub;
up-env; ; _≃_; ≃-sym; Denotation; Env) up-env; ; _≃_; ≃-sym; Denotation; Env)
open denotational.Denotational.≃-Reasoning open plfa.Denotational.≃-Reasoning
open import Data.Product using (_×_; Σ; Σ-syntax; ∃; ∃-syntax; proj₁; proj₂) open import Data.Product using (_×_; Σ; Σ-syntax; ∃; ∃-syntax; proj₁; proj₂)
renaming (_,_ to ⟨_,_⟩) renaming (_,_ to ⟨_,_⟩)

View file

@ -13,12 +13,11 @@ module plfa.Confluence where
## Imports ## Imports
\begin{code} \begin{code}
open import plfa.Substitution using (subst-commute; rename-subst-commute) open import plfa.Substitution
using (subst-commute; rename-subst-commute; Rename; Subst)
open import plfa.LambdaReduction open import plfa.LambdaReduction
using (_—→_; β; ξ₁; ξ₂; ζ; _—↠_; _—→⟨_⟩_; _[]; using (_—→_; β; ξ₁; ξ₂; ζ; _—↠_; _—→⟨_⟩_; _[];
abs-cong; appL-cong; appR-cong; —↠-trans) abs-cong; appL-cong; appR-cong; —↠-trans)
open import plfa.Denotational using (Rename)
open import plfa.Soundness using (Subst)
open import plfa.Untyped open import plfa.Untyped
using (_⊢_; _∋_; `_; _,_; ★; ƛ_; _·_; _[_]; using (_⊢_; _∋_; `_; _,_; ★; ƛ_; _·_; _[_];
rename; ext; exts; Z; S_; subst; subst-zero) rename; ext; exts; Z; S_; subst; subst-zero)

View file

@ -7,19 +7,19 @@ next : /Acknowledgements/
--- ---
\begin{code} \begin{code}
module denotational.ContextualEquivalence where module plfa.ContextualEquivalence where
\end{code} \end{code}
## Imports ## Imports
\begin{code} \begin{code}
open import plfa.Untyped using (_⊢_; ★; ∅; _,_; ƛ_) open import plfa.Untyped using (_⊢_; ★; ∅; _,_; ƛ_)
open import denotational.LambdaReduction using (_—↠_) open import plfa.LambdaReduction using (_—↠_)
open import denotational.Denotational using (; _≃_; ≃-sym; ≃-trans; _iff_) open import plfa.Denotational using (; _≃_; ≃-sym; ≃-trans; _iff_)
open import denotational.Compositional using (Ctx; plug; compositionality) open import plfa.Compositional using (Ctx; plug; compositionality)
open import denotational.Soundness using (soundness) open import plfa.Soundness using (soundness)
open import denotational.Adequacy using (adequacy) open import plfa.Adequacy using (adequacy)
open import denotational.CallByName using (_⊢_⇓_; cbn→reduce) open import plfa.CallByName using (_⊢_⇓_; cbn→reduce)
open import Data.Product using (_×_; Σ; Σ-syntax; ∃; ∃-syntax; proj₁; proj₂) open import Data.Product using (_×_; Σ; Σ-syntax; ∃; ∃-syntax; proj₁; proj₂)
renaming (_,_ to ⟨_,_⟩) renaming (_,_ to ⟨_,_⟩)

View file

@ -7,7 +7,7 @@ next : /Compositional/
--- ---
\begin{code} \begin{code}
module denotational.Denotational where module plfa.Denotational where
\end{code} \end{code}
The lambda calculus is a language about _functions_, that is, mappings The lambda calculus is a language about _functions_, that is, mappings
@ -63,7 +63,7 @@ open import Agda.Primitive using (lzero)
open import plfa.Untyped open import plfa.Untyped
using (Context; ★; _∋_; ∅; _,_; Z; S_; _⊢_; `_; _·_; ƛ_; using (Context; ★; _∋_; ∅; _,_; Z; S_; _⊢_; `_; _·_; ƛ_;
#_; twoᶜ; ext; rename; exts; subst; subst-zero; _[_]) #_; twoᶜ; ext; rename; exts; subst; subst-zero; _[_])
open import denotational.Substitution using (Rename; extensionality; rename-id) open import plfa.Substitution using (Rename; extensionality; rename-id)
open import Relation.Nullary using (¬_) open import Relation.Nullary using (¬_)
open import Relation.Nullary.Negation using (contradiction) open import Relation.Nullary.Negation using (contradiction)
open import Data.Empty using (⊥-elim) open import Data.Empty using (⊥-elim)

View file

@ -7,7 +7,7 @@ next : /Confluence/
--- ---
\begin{code} \begin{code}
module denotational.LambdaReduction where module plfa.LambdaReduction where
\end{code} \end{code}
## Imports ## Imports

View file

@ -7,7 +7,7 @@ next : /Adequacy/
--- ---
\begin{code} \begin{code}
module denotational.Soundness where module plfa.Soundness where
\end{code} \end{code}
## Imports ## Imports
@ -16,15 +16,15 @@ module denotational.Soundness where
open import plfa.Untyped open import plfa.Untyped
using (Context; _,_; _∋_; _⊢_; ★; Z; S_; `_; ƛ_; _·_; using (Context; _,_; _∋_; _⊢_; ★; Z; S_; `_; ƛ_; _·_;
subst; _[_]; subst-zero; ext; rename; exts) subst; _[_]; subst-zero; ext; rename; exts)
open import denotational.LambdaReduction open import plfa.LambdaReduction
using (_—→_; ξ₁; ξ₂; β; ζ; _—↠_; _—→⟨_⟩_; _[]) using (_—→_; ξ₁; ξ₂; β; ζ; _—↠_; _—→⟨_⟩_; _[])
open import denotational.Substitution using (Rename; Subst; ids) open import plfa.Substitution using (Rename; Subst; ids)
open import denotational.Denotational open import plfa.Denotational
using (Value; ⊥; Env; _⊢_↓_; _`,_; _⊑_; _`⊑_; `⊥; _`⊔_; init; last; init-last; using (Value; ⊥; Env; _⊢_↓_; _`,_; _⊑_; _`⊑_; `⊥; _`⊔_; init; last; init-last;
Refl⊑; Trans⊑; `Refl⊑; Env⊑; EnvConjR1⊑; EnvConjR2⊑; up-env; Refl⊑; Trans⊑; `Refl⊑; Env⊑; EnvConjR1⊑; EnvConjR2⊑; up-env;
var; ↦-elim; ↦-intro; ⊥-intro; ⊔-intro; sub; var; ↦-elim; ↦-intro; ⊥-intro; ⊔-intro; sub;
rename-pres; ; _≃_; ≃-trans) rename-pres; ; _≃_; ≃-trans)
open import denotational.Compositional using (lambda-inversion; var-inv) open import plfa.Compositional using (lambda-inversion; var-inv)
open import Relation.Binary.PropositionalEquality open import Relation.Binary.PropositionalEquality
using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app) using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app)

View file

@ -8,7 +8,7 @@ next : /LambdaReduction/
\begin{code} \begin{code}
module denotational.Substitution where module plfa.Substitution where
\end{code} \end{code}
## Imports ## Imports