From 8abb1b5601fabf5e7560d08faa6e633e60663e0a Mon Sep 17 00:00:00 2001 From: Chloe Brown Date: Thu, 29 Apr 2021 20:58:04 +0100 Subject: Finally prove that e [ μ e / zero ] ≈ μ e. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Complete proof of generator. --- src/Cfe/Expression/Properties.agda | 267 ++++++++++++++++++++++++++++--------- 1 file changed, 207 insertions(+), 60 deletions(-) (limited to 'src/Cfe/Expression/Properties.agda') diff --git a/src/Cfe/Expression/Properties.agda b/src/Cfe/Expression/Properties.agda index d994fe6..e520f7b 100644 --- a/src/Cfe/Expression/Properties.agda +++ b/src/Cfe/Expression/Properties.agda @@ -10,7 +10,6 @@ open Setoid over using () renaming (Carrier to C; _≈_ to _∼_) open import Algebra open import Cfe.Expression.Base over -open import Cfe.Function.Power open import Cfe.Language over hiding ( ≈-isPartialEquivalence; partialSetoid @@ -33,26 +32,31 @@ open import Cfe.Language over ; ∙-zeroʳ to ∙ˡ-zeroʳ ; ∙-zero to ∙ˡ-zero ) +open import Cfe.Vec.Relation.Binary.Pointwise.Inductive using (Pointwise-insert) open import Data.Empty using (⊥-elim) open import Data.Fin hiding (_+_; _≤_; _<_) +open import Data.Fin.Properties using (punchIn-punchOut) open import Data.List using (List; length; _++_) open import Data.List.Properties using (length-++) +open import Data.List.Relation.Binary.Equality.Setoid over using (_≋_) open import Data.List.Relation.Binary.Pointwise using (Pointwise-length) open import Data.Nat hiding (_≟_; _⊔_; _^_) open import Data.Nat.Induction using (<-wellFounded) open import Data.Nat.Properties hiding (_≟_) open import Data.Product open import Data.Product.Relation.Binary.Lex.Strict using (×-wellFounded) -open import Data.Sum using (_⊎_; inj₁; inj₂) -open import Data.Vec hiding (length; _++_) +open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_]′) +open import Data.Vec hiding (length; map; _++_) open import Data.Vec.Properties -open import Data.Vec.Relation.Binary.Pointwise.Inductive as PW - hiding (refl; sym; trans; setoid; lookup) +open import Data.Vec.Relation.Binary.Pointwise.Inductive as Pw + hiding (refl; sym; trans; setoid; lookup; map) +open import Function using (_∘_; _|>_; id) open import Induction.WellFounded using (WellFounded) open import Level using (_⊔_) open import Relation.Binary.Construct.On using () renaming (wellFounded to on-wellFounded) open import Relation.Binary.PropositionalEquality hiding (setoid) open import Relation.Nullary +open import Relation.Nullary.Decidable using (fromWitness) private variable @@ -71,7 +75,7 @@ private ⟦⟧-mono-env (Char _) mono = ⊆-refl ⟦⟧-mono-env (e₁ ∨ e₂) mono = ∪-mono (⟦⟧-mono-env e₁ mono) (⟦⟧-mono-env e₂ mono) ⟦⟧-mono-env (e₁ ∙ e₂) mono = ∙-mono (⟦⟧-mono-env e₁ mono) (⟦⟧-mono-env e₂ mono) -⟦⟧-mono-env (Var j) mono = PW.lookup mono j +⟦⟧-mono-env (Var j) mono = Pw.lookup mono j ⟦⟧-mono-env (μ e) mono = ⋃-mono λ A⊆B → ⟦⟧-mono-env e (A⊆B ∷ mono) ⟦⟧-cong-env : ∀ (e : Expression n) {γ γ′} → Pointwise _≈ˡ_ γ γ′ → ⟦ e ⟧ γ ≈ˡ ⟦ e ⟧ γ′ @@ -80,7 +84,7 @@ private ⟦⟧-cong-env (Char _) cong = ≈ˡ-refl ⟦⟧-cong-env (e₁ ∨ e₂) cong = ∪-cong (⟦⟧-cong-env e₁ cong) (⟦⟧-cong-env e₂ cong) ⟦⟧-cong-env (e₁ ∙ e₂) cong = ∙ˡ-cong (⟦⟧-cong-env e₁ cong) (⟦⟧-cong-env e₂ cong) -⟦⟧-cong-env (Var j) cong = PW.lookup cong j +⟦⟧-cong-env (Var j) cong = Pw.lookup cong j ⟦⟧-cong-env (μ e) cong = ⋃-cong λ A≈B → ⟦⟧-cong-env e (A≈B ∷ cong) ------------------------------------------------------------------------ @@ -128,10 +132,14 @@ setoid {n} = record { isEquivalence = ≈-isEquivalence {n} } -- Definitions infix 4 _<ₗₑₓ_ +infix 4 _<ₕₑₜ_ _<ₗₑₓ_ : REL (List C × Expression m) (List C × Expression n) _ (w , e) <ₗₑₓ (w′ , e′) = length w < length w′ ⊎ length w ≡ length w′ × e <ᵣₐₙₖ e′ +_<ₕₑₜ_ : Rel (List C × ∃[ m ] Expression m) _ +(w , _ , e) <ₕₑₜ (w′ , _ , e′) = (w , e) <ₗₑₓ (w′ , e′) + ------------------------------------------------------------------------ -- Relational properties @@ -166,6 +174,9 @@ _<ₗₑₓ_ : REL (List C × Expression m) (List C × Expression n) _ <ₗₑₓ-wellFounded : WellFounded (_<ₗₑₓ_ {n}) <ₗₑₓ-wellFounded = on-wellFounded (map₁ length) (×-wellFounded <-wellFounded <ᵣₐₙₖ-wellFounded) +<ₕₑₜ-wellFounded : WellFounded _<ₕₑₜ_ +<ₕₑₜ-wellFounded = + on-wellFounded (map length (rank ∘ proj₂)) (×-wellFounded <-wellFounded <-wellFounded) ------------------------------------------------------------------------ -- Other properties @@ -175,26 +186,20 @@ rank-∨ˡ e₁ e₂ = begin-strict rank e₁ ≤⟨ m≤m+n (rank e₁) (rank e₂) ⟩ rank e₁ + rank e₂ <⟨ n<1+n (rank e₁ + rank e₂) ⟩ rank (e₁ ∨ e₂) ∎ - where - open ≤-Reasoning + where open ≤-Reasoning rank-∨ʳ : ∀ (e₁ e₂ : Expression n) → e₂ <ᵣₐₙₖ e₁ ∨ e₂ rank-∨ʳ e₁ e₂ = begin-strict rank e₂ ≤⟨ m≤n+m (rank e₂) (rank e₁) ⟩ rank e₁ + rank e₂ <⟨ n<1+n (rank e₁ + rank e₂) ⟩ rank (e₁ ∨ e₂) ∎ - where - open ≤-Reasoning + where open ≤-Reasoning rank-∙ˡ : ∀ (e₁ e₂ : Expression n) → e₁ <ᵣₐₙₖ e₁ ∙ e₂ rank-∙ˡ e₁ _ = n<1+n (rank e₁) -lex-∙ˡ : - ∀ (e₁ e₂ : Expression n) γ → - ∀ {w} → (w∈⟦e₁∙e₂⟧ : w ∈ ⟦ e₁ ∙ e₂ ⟧ γ) → - let w₁ = proj₁ w∈⟦e₁∙e₂⟧ in - (w₁ , e₁) <ₗₑₓ (w , e₁ ∙ e₂) -lex-∙ˡ e₁ e₂ γ {w} (w₁ , w₂ , w₁∈⟦e₁⟧ , w₂∈⟦e₂⟧ , eq) with m≤n⇒m insert-lookup γ j |> ≈ˡ-reflexive |> ≈ˡ-sym |> ∈-resp-≈) + (tag (∈-resp-≈ (≈ˡ-reflexive (insert-lookup γ j A)) w∈⟦e⟧)) + ... | no i≢j = + 0 , + ∈-resp-≈ + (begin-equality + lookup (insert γ i A) j ≡˘⟨ cong (lookup (insert γ i A)) (punchIn-punchOut i≢j) ⟩ + lookup (insert γ i A) (punchIn i (punchOut i≢j)) ≡⟨ insert-punchIn γ i A (punchOut i≢j) ⟩ + lookup γ (punchOut i≢j) ≡˘⟨ insert-punchIn γ i (K 0) (punchOut i≢j) ⟩ + lookup (insert γ i (K 0)) (punchIn i (punchOut i≢j)) ≡⟨ cong (lookup (insert γ i (K 0))) (punchIn-punchOut i≢j) ⟩ + lookup (insert γ i (K 0)) j ∎) + w∈⟦e⟧ + get-tag (μ e) A K tag mono i γ {w} (n , w∈⟦e⟧) = map₂ (n ,_) (⟦e⟧′-tag n w∈⟦e⟧) + where + ⟦e⟧′ : ℕ → Language _ → Language _ + ⟦e⟧′ n A = Iter (λ X → ⟦ e ⟧ (X ∷ insert γ i A)) ∅ n + + ⟦e⟧′-monoʳ : ∀ n {X Y} → X ⊆ Y → ⟦e⟧′ n X ⊆ ⟦e⟧′ n Y + ⟦e⟧′-monoʳ n X⊆Y = + Iter-monoˡ + (⟦⟧-mono-env e ∘ (_∷ (Pointwise-insert i i {fromWitness refl} X⊆Y (Pw.refl ⊆-refl)))) + n + ⊆-refl + + ⟦e⟧′-tag : ∀ m {w} → w ∈ ⟦e⟧′ m A → ∃[ n ] w ∈ ⟦e⟧′ m (K n) + ⟦e⟧′-tag (suc m) {w} w∈⟦e⟧′ = + n₁ + n₂ , + ∈-resp-⊆ + (⟦⟧-mono-env + e + ( ⟦e⟧′-monoʳ m (mono (m≤m+n n₁ n₂)) + ∷ Pointwise-insert i i {fromWitness refl} (mono (m≤n+m n₂ n₁)) (Pw.refl ⊆-refl)) + ) + w∈⟦e⟧′₂ + where + n₁,w∈⟦e⟧′₁ : ∃[ z ] w ∈ ⟦ e ⟧ (⟦e⟧′ m (K z) ∷ insert γ i A) + n₁,w∈⟦e⟧′₁ = get-tag e (⟦e⟧′ m A) (⟦e⟧′ m ∘ K) (⟦e⟧′-tag m) (⟦e⟧′-monoʳ m ∘ mono) zero (insert γ i A) w∈⟦e⟧′ + + n₁ = proj₁ n₁,w∈⟦e⟧′₁ + w∈⟦e⟧′₁ = proj₂ n₁,w∈⟦e⟧′₁ + + n₂,w∈⟦e⟧′₂ : ∃[ z ] w ∈ ⟦ e ⟧ (⟦e⟧′ m (K n₁) ∷ insert γ i (K z)) + n₂,w∈⟦e⟧′₂ = get-tag e A K tag mono (suc i) (⟦e⟧′ m (K n₁) ∷ γ) w∈⟦e⟧′₁ + + n₂ = proj₁ n₂,w∈⟦e⟧′₂ + w∈⟦e⟧′₂ = proj₂ n₂,w∈⟦e⟧′₂ + + big-bit : ⟦ e ⟧ (⟦ μ e ⟧ γ ∷ γ) ⊆ ⟦ μ e ⟧ γ + big-bit = sub (λ w∈⟦e⟧ → + map suc id + (get-tag + e + (⟦ μ e ⟧ γ) + (Iter (⟦ e ⟧ ∘ (_∷ γ)) ∅) + id + (Iter-monoʳ (⟦⟧-mono-env e ∘ (_∷ Pw.refl ⊆-refl)) (⊆-min (⟦ e ⟧ (∅ ∷ γ)))) + zero + γ + w∈⟦e⟧)) -- cgit v1.2.3