From 0708355c7988345c98961cad087dc56eeb16ea7f Mon Sep 17 00:00:00 2001 From: Chloe Brown Date: Sat, 24 Apr 2021 15:30:30 +0100 Subject: Cleanup Derivation. --- src/Cfe/Derivation/Properties.agda | 152 ++++++++++++++----------------------- 1 file changed, 55 insertions(+), 97 deletions(-) (limited to 'src/Cfe/Derivation/Properties.agda') 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₁