summaryrefslogtreecommitdiff
path: root/src/Cfe/Derivation/Properties.agda
diff options
context:
space:
mode:
authorChloe Brown <chloe.brown.00@outlook.com>2021-04-24 15:30:30 +0100
committerChloe Brown <chloe.brown.00@outlook.com>2021-04-24 15:30:30 +0100
commit0708355c7988345c98961cad087dc56eeb16ea7f (patch)
tree76f4e4ef3f7a0eb0cf3f40d3d58e3563287044c4 /src/Cfe/Derivation/Properties.agda
parentc58866bea6ee251868d98a3da11f64030bb00aa7 (diff)
Cleanup Derivation.cleanup
Diffstat (limited to 'src/Cfe/Derivation/Properties.agda')
-rw-r--r--src/Cfe/Derivation/Properties.agda152
1 files changed, 55 insertions, 97 deletions
diff --git a/src/Cfe/Derivation/Properties.agda b/src/Cfe/Derivation/Properties.agda
index e89d9f1..d922f2a 100644
--- a/src/Cfe/Derivation/Properties.agda
+++ b/src/Cfe/Derivation/Properties.agda
@@ -6,111 +6,69 @@ module Cfe.Derivation.Properties
{c ℓ} (over : Setoid c ℓ)
where
-open Setoid over renaming (Carrier to C; _≈_ to _∼_)
+open Setoid over using () renaming (Carrier to C)
-open import Cfe.Context over hiding (_≋_)
-open import Cfe.Expression over hiding (_≋_)
-open import Cfe.Language over hiding (≤-refl; _≈_; _<_)
-open import Cfe.Language.Construct.Concatenate over using (Concat)
-open import Cfe.Language.Indexed.Construct.Iterate over
-open import Cfe.Judgement over
+open import Cfe.Context over using (∙,∙)
open import Cfe.Derivation.Base over
+open import Cfe.Expression over
+open import Cfe.Fin using (zero)
+open import Cfe.Judgement over
+open import Cfe.Language over hiding (_∙_)
open import Cfe.Type over using (_⊛_; _⊨_)
-open import Data.Bool using (T; not; true; false)
-open import Data.Empty using (⊥-elim)
-open import Data.Fin as F hiding (_<_)
-open import Data.List hiding (null)
-open import Data.List.Relation.Binary.Equality.Setoid over
-open import Data.Nat as ℕ hiding (_⊔_; _^_; _<_)
-open import Data.Nat.Properties using (≤-step; m≤m+n; m≤n+m; ≤-refl; n<1+n; module ≤-Reasoning)
-open import Data.Nat.Induction using () renaming (<-wellFounded to <ⁿ-wellFounded)
-open import Data.Product as Product
-open import Data.Product.Relation.Binary.Lex.Strict
-open import Data.Sum as Sum
-open import Data.Vec hiding (length; _++_)
-open import Data.Vec.Relation.Binary.Pointwise.Inductive
-open import Data.Vec.Relation.Binary.Pointwise.Extensional
-open import Function
+open import Data.Fin using (zero)
+open import Data.List using (List; []; length)
+open import Data.List.Relation.Binary.Pointwise using ([]; _∷_)
+open import Data.Nat.Properties using (n<1+n; module ≤-Reasoning)
+open import Data.Product using (_×_; _,_; -,_)
+open import Data.Sum using (inj₁; inj₂)
+open import Data.Vec using ([]; [_])
+open import Data.Vec.Relation.Binary.Pointwise.Inductive using ([]; _∷_)
+open import Function using (_∘_)
open import Induction.WellFounded
-open import Level
-open import Relation.Binary
-import Relation.Binary.Construct.On as On
-open import Relation.Binary.PropositionalEquality as ≡ hiding (subst₂; setoid)
-
-private
- infix 4 _<_
- _<_ : Rel (List C × Expression 0) _
- _<_ = ×-Lex _≡_ ℕ._<_ _<ᵣₐₙₖ_ on (Product.map₁ length)
+open import Level using (_⊔_)
+open import Relation.Binary.PropositionalEquality using (refl)
+import Relation.Binary.Reasoning.PartialOrder (⊆-poset {c ⊔ ℓ}) as ⊆-Reasoning
+open import Relation.Nullary using (¬_)
- <-wellFounded : WellFounded _<_
- <-wellFounded = On.wellFounded (Product.map₁ length) (×-wellFounded <ⁿ-wellFounded <ᵣₐₙₖ-wellFounded)
-
-l∈⟦e⟧⇒e⤇l : ∀ {e τ} → ∙,∙ ⊢ e ∶ τ → ∀ {l} → l ∈ ⟦ e ⟧ [] → e ⤇ l
-l∈⟦e⟧⇒e⤇l {e} {τ} ∙,∙⊢e∶τ {l} l∈⟦e⟧ = All.wfRec <-wellFounded _ Pred go (l , e) ∙,∙⊢e∶τ l∈⟦e⟧
+w∈⟦e⟧⇒e⤇w : ∀ {e τ} → ∙,∙ ⊢ e ∶ τ → ∀ {w} → w ∈ ⟦ e ⟧ [] → e ⤇ w
+w∈⟦e⟧⇒e⤇w {e = e} ctx⊢e∶τ {w} w∈⟦e⟧ = All.wfRec <ₗₑₓ-wellFounded _ Pred go (w , e) ctx⊢e∶τ w∈⟦e⟧
where
- Pred : List C × Expression 0 → Set _
- Pred (l , e) = ∀ {τ} → ∙,∙ ⊢ e ∶ τ → l ∈ ⟦ e ⟧ [] → e ⤇ l
-
- e[μe/0]<μe : ∀ {e τ} l → ∙,∙ ⊢ μ e ∶ τ → (l , e [ μ e / F.zero ]) < (l , μ e)
- e[μe/0]<μe {e} l (Fix ∙,τ⊢e∶τ)= inj₂ (≡.refl , (begin-strict
- rank (e [ μ e / F.zero ]) ≡⟨ subst-preserves-rank z≤n ∙,τ⊢e∶τ (Fix ∙,τ⊢e∶τ) ⟩
- rank e <⟨ n<1+n (rank e) ⟩
- ℕ.suc (rank e) ≡⟨⟩
- rank (μ e) ∎))
- where
- open ≤-Reasoning
+ Pred : (List C × Expression 0) → Set _
+ Pred (w , e) = ∀ {τ} → ∙,∙ ⊢ e ∶ τ → w ∈ ⟦ e ⟧ [] → e ⤇ w
- l₁++l₂≋l⇒∣l₁∣≤∣l∣ : ∀ {l₂ l} l₁ → l₁ ++ l₂ ≋ l → (length l₁ ℕ.< length l) ⊎ (length l₁ ≡ length l)
- l₁++l₂≋l⇒∣l₁∣≤∣l∣ [] [] = inj₂ ≡.refl
- l₁++l₂≋l⇒∣l₁∣≤∣l∣ [] (_ ∷ _) = inj₁ (s≤s z≤n)
- l₁++l₂≋l⇒∣l₁∣≤∣l∣ (_ ∷ l₁) (_ ∷ eq) = Sum.map s≤s (cong ℕ.suc) (l₁++l₂≋l⇒∣l₁∣≤∣l∣ l₁ eq)
-
- l₁++l₂≋l⇒∣l₂∣≤∣l∣ : ∀ {l₂ l} l₁ → l₁ ++ l₂ ≋ l → (length l₂ ℕ.< length l) ⊎ (length l₁ ≡ 0)
- l₁++l₂≋l⇒∣l₂∣≤∣l∣ [] _ = inj₂ ≡.refl
- l₁++l₂≋l⇒∣l₂∣≤∣l∣ (_ ∷ []) (_ ∷ []) = inj₁ (s≤s z≤n)
- l₁++l₂≋l⇒∣l₂∣≤∣l∣ (x ∷ []) (x∼y ∷ _ ∷ eq) = inj₁ ([ s≤s , (λ ()) ]′ (l₁++l₂≋l⇒∣l₂∣≤∣l∣ (x ∷ []) (x∼y ∷ eq)))
- l₁++l₂≋l⇒∣l₂∣≤∣l∣ (_ ∷ x ∷ l₁) (_ ∷ eq) = inj₁ ([ ≤-step , (λ ()) ]′ (l₁++l₂≋l⇒∣l₂∣≤∣l∣ (x ∷ l₁) eq))
-
- e₁<e₁∙e₂ : ∀ {l e₁} e₂ → (l∈⟦e₁∙e₂⟧ : l ∈ ⟦ e₁ ∙ e₂ ⟧ []) → (Concat.l₁ l∈⟦e₁∙e₂⟧ , e₁) < (l , e₁ ∙ e₂)
- e₁<e₁∙e₂ _ l∈⟦e₁∙e₂⟧ with l₁++l₂≋l⇒∣l₁∣≤∣l∣ (Concat.l₁ l∈⟦e₁∙e₂⟧) (Concat.eq l∈⟦e₁∙e₂⟧)
- ... | inj₁ ∣l₁∣<∣l∣ = inj₁ ∣l₁∣<∣l∣
- ... | inj₂ ∣l₁∣≡∣l∣ = inj₂ (∣l₁∣≡∣l∣ , ≤-refl)
-
- e₂<e₁∙e₂ : ∀ {l e₁ e₂ τ} → ∙,∙ ⊢ e₁ ∙ e₂ ∶ τ → (l∈⟦e₁∙e₂⟧ : l ∈ ⟦ e₁ ∙ e₂ ⟧ []) → (Concat.l₂ l∈⟦e₁∙e₂⟧ , e₂) < (l , e₁ ∙ e₂)
- e₂<e₁∙e₂ (Cat ∙,∙⊢e₁∶τ₁ _ τ₁⊛τ₂) l∈⟦e₁∙e₂⟧ with l₁++l₂≋l⇒∣l₂∣≤∣l∣ (Concat.l₁ l∈⟦e₁∙e₂⟧) (Concat.eq l∈⟦e₁∙e₂⟧)
- ... | inj₁ ∣l₂∣<∣l∣ = inj₁ ∣l₂∣<∣l∣
- ... | inj₂ ∣l₁∣≡0 with Concat.l₁ l∈⟦e₁∙e₂⟧ | Concat.l₁∈A l∈⟦e₁∙e₂⟧ | (_⊛_.τ₁.Null τ₁⊛τ₂) | _⊛_.¬n₁ τ₁⊛τ₂ | _⊨_.n⇒n (soundness ∙,∙⊢e₁∶τ₁ [] (ext (λ ()))) | ∣l₁∣≡0
- ... | [] | ε∈A | false | _ | n⇒n | refl = ⊥-elim (n⇒n ε∈A)
-
- l∈⟦e⟧ⁿ⇒l∈⟦e[μe/0]⟧ : ∀ {l} e n → l ∈ ((λ X → ⟦ e ⟧ (X ∷ [])) ^ n) (⟦ ⊥ ⟧ []) → l ∈ ⟦ e [ μ e / F.zero ] ⟧ []
- l∈⟦e⟧ⁿ⇒l∈⟦e[μe/0]⟧ e (suc n) l∈⟦e⟧ⁿ = _≤_.f
- (begin
- ((λ X → ⟦ e ⟧ (X ∷ [])) ^ (ℕ.suc n)) (⟦ ⊥ ⟧ []) ≡⟨⟩
- ⟦ e ⟧ (((λ X → ⟦ e ⟧ (X ∷ [])) ^ n) (⟦ ⊥ ⟧ []) ∷ []) ≤⟨ mono e (fⁿ≤⋃f (λ X → ⟦ e ⟧ (X ∷ [])) n ∷ []) ⟩
- ⟦ e ⟧ (⋃ (λ X → ⟦ e ⟧ (X ∷ [])) ∷ []) ≡⟨⟩
- ⟦ e ⟧ (⟦ μ e ⟧ [] ∷ []) ≈˘⟨ subst-fun e (μ e) F.zero [] ⟩
- ⟦ e [ μ e / F.zero ] ⟧ [] ∎)
- l∈⟦e⟧ⁿ
+ go : ∀ w,e → WfRec _<ₗₑₓ_ Pred w,e → Pred w,e
+ go ([] , ε) rec Eps w∈⟦e⟧ = Eps
+ go (w , Char c) rec (Char c) (c∼y ∷ []) = Char c∼y
+ go (w , μ e) rec (Fix ctx⊢e∶τ) w∈⟦e⟧ =
+ Fix (rec
+ (w , e [ μ e / zero ])
+ w,e[μe/0]<ₗₑₓw,μe
+ (subst₂ ctx⊢e∶τ zero (Fix ctx⊢e∶τ))
+ (∈-resp-⊆ ⟦μe⟧⊆⟦e[μe/0]⟧ w∈⟦e⟧))
where
- open import Relation.Binary.Reasoning.PartialOrder (poset _)
+ w,e[μe/0]<ₗₑₓw,μe : (w , e [ μ e / zero ]) <ₗₑₓ (w , μ e)
+ w,e[μe/0]<ₗₑₓw,μe = inj₂ (refl , (begin-strict
+ rank (e [ μ e / zero ]) ≡⟨ subst₂-pres-rank ctx⊢e∶τ zero (Fix ctx⊢e∶τ) ⟩
+ rank e <⟨ rank-μ e ⟩
+ rank (μ e) ∎))
+ where open ≤-Reasoning
- go : ∀ l,e → WfRec _<_ Pred l,e → Pred l,e
- go (l , e) rec Eps (lift refl) = Eps
- go (l , e) rec (Char c) (lift (c∼y ∷ [])) = Char c∼y
- go (l , μ e) rec (Fix ∙,τ⊢e∶τ) (n , l∈⟦e⟧ⁿ) =
- Fix (rec
- (l , e [ μ e / F.zero ])
- (e[μe/0]<μe l (Fix ∙,τ⊢e∶τ))
- (subst₂ z≤n ∙,τ⊢e∶τ (Fix ∙,τ⊢e∶τ))
- (l∈⟦e⟧ⁿ⇒l∈⟦e[μe/0]⟧ e n l∈⟦e⟧ⁿ))
- go (l , e₁ ∙ e₂) rec (∙,∙⊢e₁∙e₂∶τ @ (Cat ∙,∙⊢e₁∶τ₁ ∙,∙⊢e₂∶τ₂ _)) l∈⟦e⟧ =
+ ⟦μe⟧⊆⟦e[μe/0]⟧ : ⟦ μ e ⟧ [] ⊆ ⟦ e [ μ e / zero ] ⟧ []
+ ⟦μe⟧⊆⟦e[μe/0]⟧ = begin
+ ⟦ μ e ⟧ [] ≤⟨ ⋃-unroll (⟦⟧-mono-env e ∘ (_∷ [])) ⟩
+ ⟦ e ⟧ [ ⟦ μ e ⟧ [] ] ≈˘⟨ subst-cong e (μ e) zero [] ⟩
+ ⟦ e [ μ e / zero ] ⟧ [] ∎
+ where open ⊆-Reasoning
+ go (w , e₁ ∙ e₂) rec (Cat ctx⊢e₁∶τ₁ ctx⊢e₂∶τ₂ τ₁⊛τ₂) (w₁ , w₂ , w₁∈⟦e₁⟧ , w₂∈⟦e₂⟧ , eq) =
Cat
- (rec (l∈⟦e⟧.l₁ , e₁) (e₁<e₁∙e₂ e₂ l∈⟦e⟧) ∙,∙⊢e₁∶τ₁ l∈⟦e⟧.l₁∈A)
- (rec (l∈⟦e⟧.l₂ , e₂) (e₂<e₁∙e₂ ∙,∙⊢e₁∙e₂∶τ l∈⟦e⟧) ∙,∙⊢e₂∶τ₂ l∈⟦e⟧.l₂∈B)
- l∈⟦e⟧.eq
+ (rec (w₁ , e₁) (lex-∙ˡ e₁ e₂ [] (-, -, w₁∈⟦e₁⟧ , w₂∈⟦e₂⟧ , eq)) ctx⊢e₁∶τ₁ w₁∈⟦e₁⟧)
+ (rec (w₂ , e₂) (lex-∙ʳ e₁ e₂ [] ε∉⟦e₁⟧ (-, -, w₁∈⟦e₁⟧ , w₂∈⟦e₂⟧ , eq)) ctx⊢e₂∶τ₂ w₂∈⟦e₂⟧)
+ eq
where
- module l∈⟦e⟧ = Concat l∈⟦e⟧
- go (l , e₁ ∨ e₂) rec (Vee ∙,∙⊢e₁∶τ₁ _ _) (inj₁ l∈⟦e₁⟧) =
- Veeˡ (rec (l , e₁) (inj₂ (≡.refl , e₁<ᵣₐₙₖe₁∨e₂ e₁ e₂)) ∙,∙⊢e₁∶τ₁ l∈⟦e₁⟧)
- go (l , e₁ ∨ e₂) rec (Vee _ ∙,∙⊢e₂∶τ₂ _) (inj₂ l∈⟦e₂⟧) =
- Veeʳ (rec (l , e₂) (inj₂ (≡.refl , e₂<ᵣₐₙₖe₁∨e₂ e₁ e₂)) ∙,∙⊢e₂∶τ₂ l∈⟦e₂⟧)
+ open _⊛_ τ₁⊛τ₂ using (¬n₁)
+ ε∉⟦e₁⟧ : ¬ Null (⟦ e₁ ⟧ [])
+ ε∉⟦e₁⟧ = ¬n₁ ∘ _⊨_.n⇒n (soundness ctx⊢e₁∶τ₁ [] [])
+ go (w , e₁ ∨ e₂) rec (Vee ctx⊢e₁∶τ₁ ctx⊢e₂∶τ₂ τ₁#τ₂) (inj₁ w∈⟦e₁⟧) =
+ Veeˡ (rec (w , e₁) (inj₂ (refl , rank-∨ˡ e₁ e₂)) ctx⊢e₁∶τ₁ w∈⟦e₁⟧)
+ go (w , e₁ ∨ e₂) rec (Vee ctx⊢e₁∶τ₁ ctx⊢e₂∶τ₂ τ₁#τ₂) (inj₂ w∈⟦e₂⟧) =
+ Veeʳ (rec (w , e₂) (inj₂ (refl , rank-∨ʳ e₁ e₂)) ctx⊢e₂∶τ₂ w∈⟦e₂⟧)