From fff325fa2e7696bc7e14e9b7d0fd0e00ba1f9f67 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Sun, 10 Mar 2024 08:25:00 +0000 Subject: [PATCH 01/65] complete refactoring to isolate dependency on `Homogeneous` --- .../Binary/Permutation/Homogeneous.agda | 334 ++++++++++++++++-- .../Relation/Binary/Permutation/Setoid.agda | 53 +-- .../Binary/Permutation/Setoid/Properties.agda | 282 +++------------ 3 files changed, 387 insertions(+), 282 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 12015b3076..11aafd7aab 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -8,33 +8,69 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where -open import Data.List.Base using (List; _∷_) -open import Data.List.Relation.Binary.Pointwise.Base as Pointwise - using (Pointwise) -open import Data.List.Relation.Binary.Pointwise.Properties as Pointwise - using (symmetric) +open import Algebra using (IsCommutativeMonoid; Op₂) +open import Data.List.Base as List using (List; []; _∷_) +open import Data.List.Relation.Binary.Pointwise as Pointwise + using (Pointwise; _∷_) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +open import Data.List.Relation.Unary.All as All using (All; []; _∷_) +open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.Nat.Base using (ℕ; suc; _+_; _<_) +open import Data.Nat.Induction +open import Data.Nat.Properties +open import Data.Product.Base using (_,_; _×_; ∃₂) +open import Function.Base using (_∘_) open import Level using (Level; _⊔_) -open import Relation.Binary.Core using (Rel; _⇒_) +open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_) open import Relation.Binary.Bundles using (Setoid) +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Structures using (IsEquivalence) -open import Relation.Binary.Definitions using (Reflexive; Symmetric) +open import Relation.Binary.Definitions + using ( Reflexive; Symmetric; Transitive + ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) +open import Relation.Binary.PropositionalEquality.Core as ≡ + using (_≡_ ; cong) +open import Relation.Nullary.Decidable using (yes; no; does) +open import Relation.Nullary.Negation using (contradiction) +open import Relation.Unary using (Pred; Decidable) private variable - a r s : Level - A : Set a + a p r s : Level + A B : Set a + xs ys zs : List A -data Permutation {A : Set a} (R : Rel A r) : Rel (List A) (a ⊔ r) where - refl : ∀ {xs ys} → Pointwise R xs ys → Permutation R xs ys - prep : ∀ {xs ys x y} (eq : R x y) → Permutation R xs ys → Permutation R (x ∷ xs) (y ∷ ys) - swap : ∀ {xs ys x y x′ y′} (eq₁ : R x x′) (eq₂ : R y y′) → Permutation R xs ys → Permutation R (x ∷ y ∷ xs) (y′ ∷ x′ ∷ ys) - trans : ∀ {xs ys zs} → Permutation R xs ys → Permutation R ys zs → Permutation R xs zs +------------------------------------------------------------------------ +-- Definition + +module _ {A : Set a} (R : Rel A r) where + + data Permutation : Rel (List A) (a ⊔ r) where + refl : ∀ {xs ys} → Pointwise R xs ys → Permutation xs ys + prep : ∀ {xs ys x y} (eq : R x y) → Permutation xs ys → Permutation (x ∷ xs) (y ∷ ys) + swap : ∀ {xs ys x y x′ y′} (eq₁ : R x x′) (eq₂ : R y y′) → + Permutation xs ys → Permutation (x ∷ y ∷ xs) (y′ ∷ x′ ∷ ys) + trans : Transitive Permutation + +------------------------------------------------------------------------ +-- Functions over permutations + +module _ {R : Rel A r} where + + steps : Permutation R xs ys → ℕ + steps (refl _) = 1 + steps (prep _ xs↭ys) = suc (steps xs↭ys) + steps (swap _ _ xs↭ys) = suc (steps xs↭ys) + steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs ------------------------------------------------------------------------ -- The Permutation relation is an equivalence module _ {R : Rel A r} where + perm-refl : Reflexive R → Reflexive (Permutation R) + perm-refl R-refl = refl (Pointwise.refl R-refl) + sym : Symmetric R → Symmetric (Permutation R) sym R-sym (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys) sym R-sym (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (sym R-sym xs↭ys) @@ -43,7 +79,7 @@ module _ {R : Rel A r} where isEquivalence : Reflexive R → Symmetric R → IsEquivalence (Permutation R) isEquivalence R-refl R-sym = record - { refl = refl (Pointwise.refl R-refl) + { refl = perm-refl R-refl ; sym = sym R-sym ; trans = trans } @@ -53,9 +89,265 @@ module _ {R : Rel A r} where { isEquivalence = isEquivalence R-refl R-sym } -map : ∀ {R : Rel A r} {S : Rel A s} → - (R ⇒ S) → (Permutation R ⇒ Permutation S) -map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys) -map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys) -map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys) -map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) +------------------------------------------------------------------------ +-- Map + +module _ {R : Rel A r} {S : Rel A s} where + + map : (R ⇒ S) → (Permutation R ⇒ Permutation S) + map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys) + map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys) + map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys) + map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) + +------------------------------------------------------------------------ +-- Relationships to other predicates +------------------------------------------------------------------------ + +module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where + + All-resp-↭ : (All P) Respects (Permutation R) + All-resp-↭ (refl xs≋ys) pxs = Pointwise.All-resp-Pointwise resp xs≋ys pxs + All-resp-↭ (prep x≈y p) (px ∷ pxs) = resp x≈y px ∷ All-resp-↭ p pxs + All-resp-↭ (swap ≈₁ ≈₂ p) (px ∷ py ∷ pxs) = resp ≈₂ py ∷ resp ≈₁ px ∷ All-resp-↭ p pxs + All-resp-↭ (trans p₁ p₂) pxs = All-resp-↭ p₂ (All-resp-↭ p₁ pxs) + + Any-resp-↭ : (Any P) Respects (Permutation R) + Any-resp-↭ (refl xs≋ys) pxs = Pointwise.Any-resp-Pointwise resp xs≋ys pxs + Any-resp-↭ (prep x≈y p) (here px) = here (resp x≈y px) + Any-resp-↭ (prep x≈y p) (there pxs) = there (Any-resp-↭ p pxs) + Any-resp-↭ (swap x y p) (here px) = there (here (resp x px)) + Any-resp-↭ (swap x y p) (there (here px)) = here (resp y px) + Any-resp-↭ (swap x y p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) + Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) + +module _ {R : Rel A r} {S : Rel A s} + (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where + + AllPairs-resp-↭ : (AllPairs S) Respects (Permutation R) + AllPairs-resp-↭ (refl xs≋ys) pxs = + Pointwise.AllPairs-resp-Pointwise resp xs≋ys pxs + AllPairs-resp-↭ (prep x≈y p) (∼ ∷ pxs) = + All-resp-↭ rʳ p (All.map (rˡ x≈y) ∼) ∷ AllPairs-resp-↭ p pxs + AllPairs-resp-↭ (swap eq₁ eq₂ p) ((∼₁ ∷ ∼₂) ∷ ∼₃ ∷ pxs) = + (sym (rʳ eq₂ (rˡ eq₁ ∼₁)) ∷ All-resp-↭ rʳ p (All.map (rˡ eq₂) ∼₃)) ∷ + All-resp-↭ rʳ p (All.map (rˡ eq₁) ∼₂) ∷ AllPairs-resp-↭ p pxs + AllPairs-resp-↭ (trans p₁ p₂) pxs = + AllPairs-resp-↭ p₂ (AllPairs-resp-↭ p₁ pxs) + +------------------------------------------------------------------------ +-- Properties of steps, and properties of Permutation, +-- which may depend on properties of the underlying relation +------------------------------------------------------------------------ + +module _ {R : Rel A r} where + + 0 Date: Sun, 10 Mar 2024 11:23:35 +0000 Subject: [PATCH 02/65] further refactoring: `Setoid` instantiates `Homogeneous`; decouples `Pointwise`; smarter constructors --- .../Binary/Permutation/Homogeneous.agda | 52 +++++++++++++++---- .../Binary/Permutation/Propositional.agda | 4 +- .../Permutation/Propositional/Properties.agda | 25 +++++---- .../Relation/Binary/Permutation/Setoid.agda | 26 ++++------ 4 files changed, 69 insertions(+), 38 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 11aafd7aab..2948e5ce33 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -26,7 +26,7 @@ open import Relation.Binary.Bundles using (Setoid) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions - using ( Reflexive; Symmetric; Transitive + using ( Reflexive; Symmetric; Transitive; LeftTrans; RightTrans ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_ ; cong) @@ -63,13 +63,19 @@ module _ {R : Rel A r} where steps (swap _ _ xs↭ys) = suc (steps xs↭ys) steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs +-- Constructor alias + + ↭-pointwise : (Pointwise R) ⇒ Permutation R + ↭-pointwise = refl + + ------------------------------------------------------------------------ -- The Permutation relation is an equivalence module _ {R : Rel A r} where - perm-refl : Reflexive R → Reflexive (Permutation R) - perm-refl R-refl = refl (Pointwise.refl R-refl) + ↭-refl′ : Reflexive R → Reflexive (Permutation R) + ↭-refl′ R-refl = ↭-pointwise (Pointwise.refl R-refl) sym : Symmetric R → Symmetric (Permutation R) sym R-sym (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys) @@ -79,7 +85,7 @@ module _ {R : Rel A r} where isEquivalence : Reflexive R → Symmetric R → IsEquivalence (Permutation R) isEquivalence R-refl R-sym = record - { refl = perm-refl R-refl + { refl = ↭-refl′ R-refl ; sym = sym R-sym ; trans = trans } @@ -150,6 +156,15 @@ module _ {R : Rel A r} where <-≤-trans (0 Date: Sun, 10 Mar 2024 11:40:59 +0000 Subject: [PATCH 03/65] use smart transitivity --- .../Binary/Permutation/Homogeneous.agda | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 2948e5ce33..09363991a2 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -260,7 +260,7 @@ module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) where with ps′ , qs′ , eq′ , ↭′ ← helper ps qs (↭-respʳ-≋ R-trans eq xs↭ys) (rec (≡.subst (_< _) (≡.sym (steps-respʳ R-trans eq xs↭ys)) (m Date: Sun, 10 Mar 2024 11:55:07 +0000 Subject: [PATCH 04/65] refactor the `CommutativeMonoid` property --- src/Data/List/Relation/Binary/Permutation/Homogeneous.agda | 7 +++---- .../Relation/Binary/Permutation/Setoid/Properties.agda | 6 +++++- 2 files changed, 8 insertions(+), 5 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 09363991a2..57e433c358 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -8,7 +8,7 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where -open import Algebra using (IsCommutativeMonoid; Op₂) +open import Algebra.Bundles using (CommutativeMonoid) open import Data.List.Base as List using (List; []; _∷_) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; _∷_) @@ -366,9 +366,8 @@ module _ {R : Rel A r} (R-sym : Symmetric R) -- foldr of Commutative Monoid -module _ {_≈_ : Rel A r} {_∙_ : Op₂ A} {ε : A} - (isCommutativeMonoid : IsCommutativeMonoid _≈_ _∙_ ε) where - open module CM = IsCommutativeMonoid isCommutativeMonoid +module _ (commutativeMonoid : CommutativeMonoid a r) where + open module CM = CommutativeMonoid commutativeMonoid private foldr = List.foldr diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index fb5a0f705d..26c99dbbdd 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -311,5 +311,9 @@ module _ {ℓ} {R : Rel A ℓ} (R? : B.Decidable R) where module _ {_∙_ : Op₂ A} {ε : A} (isCommutativeMonoid : IsCommutativeMonoid _≈_ _∙_ ε) where + private + commutativeMonoid : CommutativeMonoid _ _ + commutativeMonoid = record { isCommutativeMonoid = isCommutativeMonoid } + foldr-commMonoid : ∀ {xs ys} → xs ↭ ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys - foldr-commMonoid = Properties.foldr-commMonoid isCommutativeMonoid + foldr-commMonoid = Properties.foldr-commMonoid commutativeMonoid From 18cb5c24794705736eb961e004157e9e8dde32ba Mon Sep 17 00:00:00 2001 From: James McKinna Date: Sun, 10 Mar 2024 12:56:27 +0000 Subject: [PATCH 05/65] refactor the `PermutationReasoning` module --- .../List/Relation/Binary/Permutation/Setoid.agda | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 771eeb6e7b..aa17e2160f 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -83,15 +83,20 @@ steps = Homogeneous.steps ------------------------------------------------------------------------ -- A reasoning API to chain permutation proofs -module PermutationReasoning where +module ↭-Reasoning (↭-isEquivalence : IsEquivalence _↭_) + (↭-pointwise : _≋_ ⇒ _↭_) + (↭-prep : ∀ x {xs ys} → xs ↭ ys → x ∷ xs ↭ x ∷ ys) + (↭-swap : ∀ x y {xs ys} → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys) + where - private module Base = ≈-Reasoning ↭-setoid + private + module Base = ≈-Reasoning (record { isEquivalence = ↭-isEquivalence }) open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) renaming (≈-go to ↭-go) - open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go ↭-sym public + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go (IsEquivalence.sym ↭-isEquivalence) public open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ (↭-go ∘′ ↭-pointwise) ≋-sym public -- Some extra combinators that allow us to skip certain elements @@ -110,3 +115,6 @@ module PermutationReasoning where syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z + +module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise ↭-prep ↭-swap + From 7749cbf4c25b796286a1528c99c3301137a86e74 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Sun, 10 Mar 2024 18:56:42 +0000 Subject: [PATCH 06/65] refactored `Propositional` through improved API for `Homogeneous` --- .../Binary/Permutation/Homogeneous.agda | 31 +++++++-- .../Binary/Permutation/Propositional.agda | 66 +++++++------------ 2 files changed, 50 insertions(+), 47 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 57e433c358..5ccc7872be 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -11,7 +11,7 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where open import Algebra.Bundles using (CommutativeMonoid) open import Data.List.Base as List using (List; []; _∷_) open import Data.List.Relation.Binary.Pointwise as Pointwise - using (Pointwise; _∷_) + using (Pointwise; []; _∷_) open import Data.List.Relation.Unary.Any as Any using (Any; here; there) open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) @@ -68,6 +68,16 @@ module _ {R : Rel A r} where ↭-pointwise : (Pointwise R) ⇒ Permutation R ↭-pointwise = refl +-- Smart eliminators + + ↭-[]-invˡ : Permutation R [] xs → xs ≡ [] + ↭-[]-invˡ (refl []) = ≡.refl + ↭-[]-invˡ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invˡ xs↭ys = ↭-[]-invˡ ys↭zs + + ↭-[]-invʳ : Permutation R xs [] → xs ≡ [] + ↭-[]-invʳ (refl []) = ≡.refl + ↭-[]-invʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ↭-[]-invʳ xs↭ys + ------------------------------------------------------------------------ -- The Permutation relation is an equivalence @@ -95,6 +105,20 @@ module _ {R : Rel A r} where { isEquivalence = isEquivalence R-refl R-sym } +------------------------------------------------------------------------ +-- A smart version of trans that pushes `refl`s to the leaves (see #1113). + +module _ {R : Rel A r} + (↭-transˡ-≋ : LeftTrans (Pointwise R) (Permutation R)) + (↭-transʳ-≋ : RightTrans (Permutation R) (Pointwise R)) + where + + ↭-trans′ : Transitive (Permutation R) + ↭-trans′ (refl xs≋ys) ys↭zs = ↭-transˡ-≋ xs≋ys ys↭zs + ↭-trans′ xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs + ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs + + ------------------------------------------------------------------------ -- Map @@ -196,11 +220,8 @@ module _ {R : Rel A r} (R-trans : Transitive R) where ↭-transʳ-≋ (swap eq₁ eq₂ xs↭ys) (x≈w ∷ y≈z ∷ ys≋zs) = swap (R-trans eq₁ y≈z) (R-trans eq₂ x≈w) (↭-transʳ-≋ xs↭ys ys≋zs) ↭-transʳ-≋ (trans xs↭ws ws↭ys) ys≋zs = trans xs↭ws (↭-transʳ-≋ ws↭ys ys≋zs) --- A smart version of trans that pushes `refl`s to the leaves (see #1113). ↭-trans : Transitive (Permutation R) - ↭-trans (refl xs≋ys) ys↭zs = ↭-transˡ-≋ xs≋ys ys↭zs - ↭-trans xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs - ↭-trans xs↭ys ys↭zs = trans xs↭ys ys↭zs + ↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ module _ {R : Rel A r} (R-sym : Symmetric R) (R-trans : Transitive R) where diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 6a6b898bb5..232533076b 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -13,11 +13,17 @@ open import Data.List.Base using (List; []; _∷_) open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) -open import Relation.Binary.Definitions using (Reflexive; Transitive) -open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) +open import Relation.Binary.Definitions + using (Reflexive; Transitive; LeftTrans; RightTrans) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl; setoid) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Reasoning.Syntax +open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise) +import Data.List.Relation.Binary.Permutation.Setoid as Permutation +import Data.List.Relation.Binary.Permutation.Homogeneous as Properties + + ------------------------------------------------------------------------ -- An inductive definition of permutation @@ -26,7 +32,7 @@ open import Relation.Binary.Reasoning.Syntax -- adding in a bunch of trivial `_≡_` proofs to the constructors which -- a) adds noise and b) prevents easy access to the variables `x`, `y`. -- This may be changed in future when a better solution is found. - +{- infix 3 _↭_ data _↭_ : Rel (List A) a where @@ -34,31 +40,32 @@ data _↭_ : Rel (List A) a where prep : ∀ {xs ys} x → xs ↭ ys → x ∷ xs ↭ x ∷ ys swap : ∀ {xs ys} x y → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys trans : ∀ {xs ys zs} → xs ↭ ys → ys ↭ zs → xs ↭ zs +-} +open module ↭ = Permutation (setoid A) public + using (_↭_; ↭-refl; ↭-sym; ↭-prep; ↭-swap; module ↭-Reasoning) ------------------------------------------------------------------------ -- _↭_ is an equivalence ↭-reflexive : _≡_ ⇒ _↭_ -↭-reflexive refl = refl - -↭-refl : Reflexive _↭_ -↭-refl = refl +↭-reflexive refl = ↭-refl -↭-sym : ∀ {xs ys} → xs ↭ ys → ys ↭ xs -↭-sym refl = refl -↭-sym (prep x xs↭ys) = prep x (↭-sym xs↭ys) -↭-sym (swap x y xs↭ys) = swap y x (↭-sym xs↭ys) -↭-sym (trans xs↭ys ys↭zs) = trans (↭-sym ys↭zs) (↭-sym xs↭ys) +↭-pointwise : (Pointwise _≡_) ⇒ _↭_ +↭-pointwise xs≋ys = ↭-reflexive (Pointwise.Pointwise-≡⇒≡ xs≋ys) -- A smart version of trans that avoids unnecessary `refl`s (see #1113). ↭-trans : Transitive _↭_ -↭-trans refl ρ₂ = ρ₂ -↭-trans ρ₁ refl = ρ₁ -↭-trans ρ₁ ρ₂ = trans ρ₁ ρ₂ +↭-trans = Properties.↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ + where + ↭-transˡ-≋ : LeftTrans (Pointwise _≡_) _↭_ + ↭-transˡ-≋ xs≋ys ys↭zs with refl ← Pointwise.Pointwise-≡⇒≡ xs≋ys = ys↭zs + + ↭-transʳ-≋ : RightTrans _↭_ (Pointwise _≡_) + ↭-transʳ-≋ xs↭ys ys≋zs with refl ← Pointwise.Pointwise-≡⇒≡ ys≋zs = xs↭ys ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record - { refl = refl + { refl = ↭-refl ; sym = ↭-sym ; trans = ↭-trans } @@ -72,29 +79,4 @@ data _↭_ : Rel (List A) a where -- A reasoning API to chain permutation proofs and allow "zooming in" -- to localised reasoning. -module PermutationReasoning where - - private module Base = ≈-Reasoning ↭-setoid - - open Base public - hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) - renaming (≈-go to ↭-go) - - open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go ↭-sym public - - -- Some extra combinators that allow us to skip certain elements - - infixr 2 step-swap step-prep - - -- Skip reasoning on the first element - step-prep : ∀ x xs {ys zs : List A} → (x ∷ ys) IsRelatedTo zs → - xs ↭ ys → (x ∷ xs) IsRelatedTo zs - step-prep x xs rel xs↭ys = relTo (trans (prep x xs↭ys) (begin rel)) - - -- Skip reasoning about the first two elements - step-swap : ∀ x y xs {ys zs : List A} → (y ∷ x ∷ ys) IsRelatedTo zs → - xs ↭ ys → (x ∷ y ∷ xs) IsRelatedTo zs - step-swap x y xs rel xs↭ys = relTo (trans (swap x y xs↭ys) (begin rel)) - - syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z - syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z +module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise From de26a1e7a3df8db2150abc4c1e94ff7eee88ec3f Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 05:53:47 +0000 Subject: [PATCH 07/65] tweak imports --- src/Data/List/Relation/Binary/Equality/Propositional.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Binary/Equality/Propositional.agda b/src/Data/List/Relation/Binary/Equality/Propositional.agda index ffae515086..553348a920 100644 --- a/src/Data/List/Relation/Binary/Equality/Propositional.agda +++ b/src/Data/List/Relation/Binary/Equality/Propositional.agda @@ -14,7 +14,7 @@ open import Relation.Binary.Core using (_⇒_) module Data.List.Relation.Binary.Equality.Propositional {a} {A : Set a} where -open import Data.List.Base +import Data.List.Base as List import Data.List.Relation.Binary.Equality.Setoid as SetoidEquality open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong) import Relation.Binary.PropositionalEquality.Properties as ≡ @@ -29,7 +29,7 @@ open SetoidEquality (≡.setoid A) public ≋⇒≡ : _≋_ ⇒ _≡_ ≋⇒≡ [] = refl -≋⇒≡ (refl ∷ xs≈ys) = cong (_ ∷_) (≋⇒≡ xs≈ys) +≋⇒≡ (refl ∷ xs≈ys) = cong (_ List.∷_) (≋⇒≡ xs≈ys) ≡⇒≋ : _≡_ ⇒ _≋_ ≡⇒≋ refl = ≋-refl From a2361802c82bf048f2aef6a51e3de43bd5e5af4d Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 05:54:26 +0000 Subject: [PATCH 08/65] refactor `setoid` and `Reasoning` --- .../Relation/Binary/Permutation/Setoid.agda | 33 +++++++++++-------- 1 file changed, 20 insertions(+), 13 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index aa17e2160f..14210b09ab 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -11,7 +11,7 @@ open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions - using (Reflexive; Symmetric; Transitive) + using (Reflexive; Symmetric; Transitive; LeftTrans; RightTrans) open import Relation.Binary.Reasoning.Syntax module Data.List.Relation.Binary.Permutation.Setoid @@ -38,7 +38,7 @@ _↭_ : Rel (List A) (a ⊔ ℓ) _↭_ = Homogeneous.Permutation _≈_ ------------------------------------------------------------------------ --- Constructor aliases +-- Smart constructor aliases -- These provide aliases for the `Homogeneous` smart constructors, in -- particular for `swap` and `prep` when the elements being swapped or @@ -53,11 +53,11 @@ _↭_ = Homogeneous.Permutation _≈_ ↭-swap : ∀ x y {xs ys} → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys ↭-swap = Homogeneous.↭-swap ≈-refl -↭-trans : Transitive _↭_ -↭-trans = Homogeneous.↭-trans ≈-trans +↭-trans′ : LeftTrans _≋_ _↭_ → RightTrans _↭_ _≋_ → Transitive _↭_ +↭-trans′ = Homogeneous.↭-trans′ ------------------------------------------------------------------------ --- Functions over permutations +-- Functions over permutations (retained for legacy) steps : ∀ {xs ys} → xs ↭ ys → ℕ steps = Homogeneous.steps @@ -74,29 +74,31 @@ steps = Homogeneous.steps ↭-sym : Symmetric _↭_ ↭-sym = Homogeneous.sym ≈-sym +↭-trans : Transitive _↭_ +↭-trans = Homogeneous.↭-trans ≈-trans + ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record { refl = ↭-refl ; sym = ↭-sym ; trans = ↭-trans } -↭-setoid : Setoid _ _ -↭-setoid = record { isEquivalence = ↭-isEquivalence } - ------------------------------------------------------------------------ -- A reasoning API to chain permutation proofs module ↭-Reasoning (↭-isEquivalence : IsEquivalence _↭_) (↭-pointwise : _≋_ ⇒ _↭_) - (↭-prep : ∀ x {xs ys} → xs ↭ ys → x ∷ xs ↭ x ∷ ys) - (↭-swap : ∀ x y {xs ys} → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys) where + ↭-setoid : Setoid _ _ + ↭-setoid = record { isEquivalence = ↭-isEquivalence } + private - module Base = ≈-Reasoning (record { isEquivalence = ↭-isEquivalence }) + open IsEquivalence ↭-isEquivalence using () renaming (sym to ↭-sym′) + module Base = ≈-Reasoning ↭-setoid open Base public hiding (step-≈; step-≈˘; step-≈-⟩; step-≈-⟨) renaming (≈-go to ↭-go) - open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go (IsEquivalence.sym ↭-isEquivalence) public + open ↭-syntax _IsRelatedTo_ _IsRelatedTo_ ↭-go ↭-sym′ public open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ (↭-go ∘′ ↭-pointwise) ≋-sym public -- Some extra combinators that allow us to skip certain elements @@ -116,5 +118,10 @@ module ↭-Reasoning (↭-isEquivalence : IsEquivalence _↭_) syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z -module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise ↭-prep ↭-swap +module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise + +------------------------------------------------------------------------ +-- Bundle export + +open PermutationReasoning public using (↭-setoid) From 05fde1e7ad09f6d65fa846772f52208dbad557e5 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 05:55:21 +0000 Subject: [PATCH 09/65] remove last explicit dependency on `Homogeneous` --- .../Binary/Permutation/Propositional.agda | 40 ++++++++++--------- 1 file changed, 21 insertions(+), 19 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 232533076b..90b05d4efc 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -10,18 +10,15 @@ module Data.List.Relation.Binary.Permutation.Propositional {a} {A : Set a} where open import Data.List.Base using (List; []; _∷_) +open import Data.List.Relation.Binary.Equality.Propositional using (_≋_; ≋⇒≡) open import Relation.Binary.Core using (Rel; _⇒_) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Structures using (IsEquivalence) open import Relation.Binary.Definitions using (Reflexive; Transitive; LeftTrans; RightTrans) -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl; setoid) -import Relation.Binary.Reasoning.Setoid as ≈-Reasoning -open import Relation.Binary.Reasoning.Syntax +open import Relation.Binary.PropositionalEquality using (_≡_; refl; setoid) -open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise) import Data.List.Relation.Binary.Permutation.Setoid as Permutation -import Data.List.Relation.Binary.Permutation.Homogeneous as Properties ------------------------------------------------------------------------ @@ -42,7 +39,10 @@ data _↭_ : Rel (List A) a where trans : ∀ {xs ys zs} → xs ↭ ys → ys ↭ zs → xs ↭ zs -} open module ↭ = Permutation (setoid A) public - using (_↭_; ↭-refl; ↭-sym; ↭-prep; ↭-swap; module ↭-Reasoning) + hiding ( ↭-reflexive; ↭-pointwise; ↭-trans + ; ↭-isEquivalence; ↭-setoid; module PermutationReasoning + ) + ------------------------------------------------------------------------ -- _↭_ is an equivalence @@ -50,18 +50,18 @@ open module ↭ = Permutation (setoid A) public ↭-reflexive : _≡_ ⇒ _↭_ ↭-reflexive refl = ↭-refl -↭-pointwise : (Pointwise _≡_) ⇒ _↭_ -↭-pointwise xs≋ys = ↭-reflexive (Pointwise.Pointwise-≡⇒≡ xs≋ys) +↭-pointwise : _≋_ ⇒ _↭_ +↭-pointwise xs≋ys = ↭-reflexive (≋⇒≡ xs≋ys) -- A smart version of trans that avoids unnecessary `refl`s (see #1113). ↭-trans : Transitive _↭_ -↭-trans = Properties.↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ +↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ where - ↭-transˡ-≋ : LeftTrans (Pointwise _≡_) _↭_ - ↭-transˡ-≋ xs≋ys ys↭zs with refl ← Pointwise.Pointwise-≡⇒≡ xs≋ys = ys↭zs + ↭-transˡ-≋ : LeftTrans _≋_ _↭_ + ↭-transˡ-≋ xs≋ys ys↭zs with refl ← ≋⇒≡ xs≋ys = ys↭zs - ↭-transʳ-≋ : RightTrans _↭_ (Pointwise _≡_) - ↭-transʳ-≋ xs↭ys ys≋zs with refl ← Pointwise.Pointwise-≡⇒≡ ys≋zs = xs↭ys + ↭-transʳ-≋ : RightTrans _↭_ _≋_ + ↭-transʳ-≋ xs↭ys ys≋zs with refl ← ≋⇒≡ ys≋zs = xs↭ys ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record @@ -70,13 +70,15 @@ open module ↭ = Permutation (setoid A) public ; trans = ↭-trans } -↭-setoid : Setoid _ _ -↭-setoid = record - { isEquivalence = ↭-isEquivalence - } - ------------------------------------------------------------------------ -- A reasoning API to chain permutation proofs and allow "zooming in" --- to localised reasoning. +-- to localised reasoning. For details of the axiomatisation, and +-- in particular the refactored dependencies, +-- see `Data.List.Relation.Binary.Permutation.Setoid` module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise + +------------------------------------------------------------------------ +-- Bundle export + +open PermutationReasoning public using (↭-setoid) From d4e86d93a3c34b61b4f85a91d29169c3d7d8723c Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 09:57:05 +0000 Subject: [PATCH 10/65] added length preservation --- src/Data/List/Relation/Binary/Pointwise/Properties.agda | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/Pointwise/Properties.agda b/src/Data/List/Relation/Binary/Pointwise/Properties.agda index 1761dd3cda..d203f915b2 100644 --- a/src/Data/List/Relation/Binary/Pointwise/Properties.agda +++ b/src/Data/List/Relation/Binary/Pointwise/Properties.agda @@ -9,7 +9,7 @@ module Data.List.Relation.Binary.Pointwise.Properties where open import Data.Product.Base using (_,_; uncurry) -open import Data.List.Base using (List; []; _∷_) +open import Data.List.Base as List using (List; []; _∷_) open import Level open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions @@ -25,6 +25,8 @@ private A : Set a B : Set b R S T : REL A B ℓ + xs : List A + ys : List B ------------------------------------------------------------------------ -- Relational properties From 629d3f3771fb682cd5e2d774383e8f885f4a58c6 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 09:57:47 +0000 Subject: [PATCH 11/65] added length preservation --- .../Binary/Permutation/Homogeneous.agda | 42 +++++++- .../Permutation/Propositional/Properties.agda | 95 ++++++++----------- .../Binary/Permutation/Setoid/Properties.agda | 63 ++++++++---- 3 files changed, 122 insertions(+), 78 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 5ccc7872be..e9aff5ba28 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -9,7 +9,7 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where open import Algebra.Bundles using (CommutativeMonoid) -open import Data.List.Base as List using (List; []; _∷_) +open import Data.List.Base as List using (List; []; _∷_; [_]) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) open import Data.List.Relation.Unary.Any as Any using (Any; here; there) @@ -18,7 +18,7 @@ open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) open import Data.Nat.Base using (ℕ; suc; _+_; _<_) open import Data.Nat.Induction open import Data.Nat.Properties -open import Data.Product.Base using (_,_; _×_; ∃₂) +open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) open import Function.Base using (_∘_) open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_) @@ -31,7 +31,7 @@ open import Relation.Binary.Definitions open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_ ; cong) open import Relation.Nullary.Decidable using (yes; no; does) -open import Relation.Nullary.Negation using (contradiction) +open import Relation.Nullary.Negation using (¬_; contradiction) open import Relation.Unary using (Pred; Decidable) private @@ -39,6 +39,8 @@ private a p r s : Level A B : Set a xs ys zs : List A + x y z v w : A + ------------------------------------------------------------------------ -- Definition @@ -68,7 +70,7 @@ module _ {R : Rel A r} where ↭-pointwise : (Pointwise R) ⇒ Permutation R ↭-pointwise = refl --- Smart eliminators +-- Smart inversions ↭-[]-invˡ : Permutation R [] xs → xs ≡ [] ↭-[]-invˡ (refl []) = ≡.refl @@ -78,6 +80,9 @@ module _ {R : Rel A r} where ↭-[]-invʳ (refl []) = ≡.refl ↭-[]-invʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ↭-[]-invʳ xs↭ys + ¬x∷xs↭[] : ¬ (Permutation R (x ∷ xs) []) + ¬x∷xs↭[] (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ¬x∷xs↭[] xs↭ys + ------------------------------------------------------------------------ -- The Permutation relation is an equivalence @@ -223,6 +228,17 @@ module _ {R : Rel A r} (R-trans : Transitive R) where ↭-trans : Transitive (Permutation R) ↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ +-- Smart inversion + + ↭-singleton⁻¹ : Permutation R xs [ x ] → ∃ λ y → xs ≡ [ y ] × R y x + ↭-singleton⁻¹ (refl (yRx ∷ [])) = _ , ≡.refl , yRx + ↭-singleton⁻¹ (prep yRx p) + with ≡.refl ← ↭-[]-invʳ p = _ , ≡.refl , yRx + ↭-singleton⁻¹ (trans r s) + with _ , ≡.refl , yRx ← ↭-singleton⁻¹ s + with _ , ≡.refl , zRy ← ↭-singleton⁻¹ r + = _ , ≡.refl , R-trans zRy yRx + module _ {R : Rel A r} (R-sym : Symmetric R) (R-trans : Transitive R) where @@ -320,6 +336,17 @@ module _ {R : Rel A r} -- Properties of List functions ------------------------------------------------------------------------ + +-- length + +module _ {R : Rel A r} where + + ↭-length : Permutation R xs ys → List.length xs ≡ List.length ys + ↭-length (refl xs≋ys) = Pointwise.Pointwise-length xs≋ys + ↭-length (prep _ xs↭ys) = cong suc (↭-length xs↭ys) + ↭-length (swap _ _ xs↭ys) = cong (suc ∘ suc) (↭-length xs↭ys) + ↭-length (trans xs↭ys ys↭zs) = ≡.trans (↭-length xs↭ys) (↭-length ys↭zs) + -- map module _ {R : Rel A r} {S : Rel B s} {f} (pres : f Preserves R ⟶ S) where @@ -330,7 +357,12 @@ module _ {R : Rel A r} {S : Rel B s} {f} (pres : f Preserves R ⟶ S) where map⁺ (prep x xs↭ys) = prep (pres x) (map⁺ xs↭ys) map⁺ (swap x y xs↭ys) = swap (pres x) (pres y) (map⁺ xs↭ys) map⁺ (trans xs↭ys ys↭zs) = trans (map⁺ xs↭ys) (map⁺ ys↭zs) - +{- + -- permutations preserve 'being a mapped list' + ↭-map-inv : ∀ {zs : List B} → Permutation S (List.map f xs) zs → + ∃ λ ys → zs ≡ List.map f ys × Permutation R xs ys + ↭-map-inv p = {!p!} +-} ------------------------------------------------------------------------ -- Properties of List functions which depend on diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 7a01e59895..f4a9eeacf1 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -6,7 +6,7 @@ {-# OPTIONS --cubical-compatible --safe #-} -module Data.List.Relation.Binary.Permutation.Propositional.Properties where +module Data.List.Relation.Binary.Permutation.Propositional.Properties {a} {A : Set a} where open import Algebra.Bundles open import Algebra.Definitions @@ -14,8 +14,7 @@ open import Algebra.Structures open import Data.Bool.Base using (Bool; true; false) open import Data.Nat using (suc) open import Data.Product.Base using (-,_; proj₂) -open import Data.List.Base as List -open import Data.List.Relation.Binary.Permutation.Propositional +open import Data.List.Base as List using (List; []; _∷_; [_]; _++_) open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.All using (All; []; _∷_) open import Data.List.Membership.Propositional @@ -27,65 +26,44 @@ open import Level using (Level) open import Relation.Unary using (Pred) open import Relation.Binary.Core using (Rel; _Preserves₂_⟶_⟶_) open import Relation.Binary.Definitions using (_Respects_; Decidable) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_ ; refl ; cong; cong₂; _≢_) +open import Relation.Binary.PropositionalEquality as ≡ + using (_≡_ ; refl ; cong; cong₂; _≢_; setoid) open import Relation.Nullary -open PermutationReasoning +import Data.List.Relation.Binary.Permutation.Propositional as Propositional +import Data.List.Relation.Binary.Permutation.Setoid.Properties as Permutation +import Data.List.Relation.Binary.Permutation.Homogeneous as Properties private variable - a b p : Level - A : Set a + b p : Level B : Set b + x y z : A + xs ys zs : List A ------------------------------------------------------------------------- --- Permutations of empty and singleton lists +open Propositional {A = A} +open Permutation (setoid A) public -↭-empty-inv : ∀ {xs : List A} → xs ↭ [] → xs ≡ [] -↭-empty-inv refl = refl -↭-empty-inv (trans p q) with refl ← ↭-empty-inv q = ↭-empty-inv p -¬x∷xs↭[] : ∀ {x} {xs : List A} → ¬ ((x ∷ xs) ↭ []) -¬x∷xs↭[] (trans s₁ s₂) with ↭-empty-inv s₂ -... | refl = ¬x∷xs↭[] s₁ +------------------------------------------------------------------------ +-- Permutations of empty and singleton lists -↭-singleton-inv : ∀ {x} {xs : List A} → xs ↭ [ x ] → xs ≡ [ x ] -↭-singleton-inv refl = refl -↭-singleton-inv (prep _ ρ) with refl ← ↭-empty-inv ρ = refl -↭-singleton-inv (trans ρ₁ ρ₂) with refl ← ↭-singleton-inv ρ₂ = ↭-singleton-inv ρ₁ +↭-singleton-inv : xs ↭ [ x ] → xs ≡ [ x ] +↭-singleton-inv ρ with _ , refl , refl ← ↭-singleton⁻¹ ρ = refl ------------------------------------------------------------------------ -- sym - +{- ↭-sym-involutive : ∀ {xs ys : List A} (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p ↭-sym-involutive refl = refl ↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭) ↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭) ↭-sym-involutive (trans ↭₁ ↭₂) = cong₂ trans (↭-sym-involutive ↭₁) (↭-sym-involutive ↭₂) - +-} ------------------------------------------------------------------------ -- Relationships to other predicates - -All-resp-↭ : ∀ {P : Pred A p} → (All P) Respects _↭_ -All-resp-↭ refl wit = wit -All-resp-↭ (prep x p) (px ∷ wit) = px ∷ All-resp-↭ p wit -All-resp-↭ (swap x y p) (px ∷ py ∷ wit) = py ∷ px ∷ All-resp-↭ p wit -All-resp-↭ (trans p₁ p₂) wit = All-resp-↭ p₂ (All-resp-↭ p₁ wit) - -Any-resp-↭ : ∀ {P : Pred A p} → (Any P) Respects _↭_ -Any-resp-↭ refl wit = wit -Any-resp-↭ (prep x p) (here px) = here px -Any-resp-↭ (prep x p) (there wit) = there (Any-resp-↭ p wit) -Any-resp-↭ (swap x y p) (here px) = there (here px) -Any-resp-↭ (swap x y p) (there (here px)) = here px -Any-resp-↭ (swap x y p) (there (there wit)) = there (there (Any-resp-↭ p wit)) -Any-resp-↭ (trans p p₁) wit = Any-resp-↭ p₁ (Any-resp-↭ p wit) - -∈-resp-↭ : ∀ {x : A} → (x ∈_) Respects _↭_ -∈-resp-↭ = Any-resp-↭ - +{- Any-resp-[σ⁻¹∘σ] : {xs ys : List A} {P : Pred A p} → (σ : xs ↭ ys) → (ix : Any P xs) → @@ -110,10 +88,10 @@ Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) (ix : x ∈ xs) → ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix ∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] - +-} ------------------------------------------------------------------------ -- map - +{- module _ (f : A → B) where map⁺ : ∀ {xs ys} → xs ↭ ys → map f xs ↭ map f ys @@ -123,7 +101,7 @@ module _ (f : A → B) where map⁺ (trans p₁ p₂) = trans (map⁺ p₁) (map⁺ p₂) -- permutations preserve 'being a mapped list' - ↭-map-inv : ∀ {xs ys} → map f xs ↭ ys → ∃ λ ys′ → ys ≡ map f ys′ × xs ↭ ys′ + ↭-map-inv : ∀ {xs zs} → map f xs ↭ zs → ∃ λ ys → zs ≡ map f ys × xs ↭ ys ↭-map-inv {[]} ρ = -, ↭-empty-inv (↭-sym ρ) , ↭-refl ↭-map-inv {x ∷ []} ρ = -, ↭-singleton-inv (↭-sym ρ) , ↭-refl ↭-map-inv {_ ∷ _ ∷ _} refl = -, refl , ↭-refl @@ -131,16 +109,8 @@ module _ (f : A → B) where ↭-map-inv {_ ∷ _ ∷ _} (swap _ _ ρ) with _ , refl , ρ′ ← ↭-map-inv ρ = -, refl , swap _ _ ρ′ ↭-map-inv {_ ∷ _ ∷ _} (trans ρ₁ ρ₂) with _ , refl , ρ₃ ← ↭-map-inv ρ₁ with _ , refl , ρ₄ ← ↭-map-inv ρ₂ = -, refl , trans ρ₃ ρ₄ - ------------------------------------------------------------------------- --- length - -↭-length : ∀ {xs ys : List A} → xs ↭ ys → length xs ≡ length ys -↭-length refl = refl -↭-length (prep x lr) = cong suc (↭-length lr) -↭-length (swap x y lr) = cong (suc ∘ suc) (↭-length lr) -↭-length (trans lr₁ lr₂) = ≡.trans (↭-length lr₁) (↭-length lr₂) - +-} +{- ------------------------------------------------------------------------ -- _++_ @@ -356,3 +326,20 @@ module _ {ℓ} {R : Rel A ℓ} (R? : Decidable R) where (x ∷ xs) ++ y ∷ ys ≡⟨ List.++-assoc [ x ] xs (y ∷ ys) ⟨ x ∷ xs ++ y ∷ ys ∎ where open PermutationReasoning + +-} +------------------------------------------------------------------------ +-- DEPRECATED NAMES +------------------------------------------------------------------------ +-- Please use the new names as continuing support for the old names is +-- not guaranteed. + +-- Version 2.1 + +↭-empty-inv = ↭-[]-invʳ +{-# WARNING_ON_USAGE ↭-empty-inv +"Warning: ↭-empty-inv was deprecated in v2.1. +Please use ↭-[]-invʳ instead." +#-} + + diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 26c99dbbdd..86940bd619 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -14,9 +14,9 @@ module Data.List.Relation.Binary.Permutation.Setoid.Properties open import Algebra open import Data.Bool.Base using (true; false) -open import Data.List.Base as List hiding (head; tail) +open import Data.List.Base open import Data.List.Relation.Binary.Pointwise as Pointwise - using (Pointwise; head; tail) + using (Pointwise) import Data.List.Relation.Binary.Equality.Setoid as Equality import Data.List.Relation.Binary.Permutation.Setoid as Permutation import Data.List.Relation.Binary.Permutation.Homogeneous as Properties @@ -25,7 +25,6 @@ open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) import Data.List.Relation.Unary.Unique.Setoid as Unique import Data.List.Membership.Setoid as Membership -open import Data.List.Membership.Setoid.Properties using (∈-∃++; ∈-insert) import Data.List.Properties as List open import Data.Nat hiding (_⊔_) open import Data.Product.Base using (_,_; _×_; ∃; ∃₂; proj₁; proj₂) @@ -39,13 +38,9 @@ open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_ ; refl; sym; cong; cong₂; subst; _≢_) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Nullary.Decidable using (yes; no; does) -open import Relation.Nullary.Negation using (contradiction) +open import Relation.Nullary.Negation using (¬_; contradiction) open import Relation.Unary using (Pred; Decidable) -private - variable - b p r : Level - open Setoid S using (_≈_) renaming (Carrier to A; refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) @@ -55,6 +50,11 @@ open Unique S using (Unique) open module ≋ = Equality S using (_≋_; []; _∷_; ≋-refl; ≋-sym; ≋-trans) +private + variable + b p r : Level + xs ys zs : List A + x y z v w : A ------------------------------------------------------------------------ -- Relationships to other predicates @@ -107,6 +107,27 @@ steps-respʳ = Properties.steps-respʳ ≈-trans -- Properties of list functions ------------------------------------------------------------------------ +------------------------------------------------------------------------ +-- Permutations of empty and singleton lists + +↭-[]-invˡ : [] ↭ xs → xs ≡ [] +↭-[]-invˡ = Properties.↭-[]-invˡ + +↭-[]-invʳ : xs ↭ [] → xs ≡ [] +↭-[]-invʳ = Properties.↭-[]-invʳ + +¬x∷xs↭[] : ¬ ((x ∷ xs) ↭ []) +¬x∷xs↭[] = Properties.¬x∷xs↭[] + +↭-singleton⁻¹ : xs ↭ [ x ] → ∃ λ y → xs ≡ [ y ] × y ≈ x +↭-singleton⁻¹ = Properties.↭-singleton⁻¹ ≈-trans + +------------------------------------------------------------------------ +-- length + +↭-length : xs ↭ ys → length xs ≡ length ys +↭-length = Properties.↭-length + ------------------------------------------------------------------------ -- map @@ -116,23 +137,27 @@ module _ (T : Setoid b ℓ) where open Permutation T using () renaming (_↭_ to _↭′_) map⁺ : ∀ {f} → f Preserves _≈_ ⟶ _≈′_ → - ∀ {xs ys} → xs ↭ ys → map f xs ↭′ map f ys - map⁺ = Properties.map⁺ - + xs ↭ ys → map f xs ↭′ map f ys + map⁺ pres = Properties.map⁺ pres +{- + -- permutations preserve 'being a mapped list' + ↭-map-inv : map f xs ↭ zs → ∃ λ ys → zs ≡ map f ys × xs ↭ ys + ↭-map-inv p = +-} ------------------------------------------------------------------------ -- _++_ -shift : ∀ {v w} → v ≈ w → (xs ys : List A) → xs ++ [ v ] ++ ys ↭ w ∷ xs ++ ys +shift : v ≈ w → ∀ xs ys → xs ++ [ v ] ++ ys ↭ w ∷ xs ++ ys shift v≈w xs ys = Properties.shift ≈-refl ≈-sym ≈-trans v≈w xs {ys} -↭-shift : ∀ {v} (xs ys : List A) → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys +↭-shift : ∀ xs ys → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys ↭-shift = shift ≈-refl -++⁺ˡ : ∀ xs {ys zs : List A} → ys ↭ zs → xs ++ ys ↭ xs ++ zs +++⁺ˡ : ∀ xs {ys zs} → ys ↭ zs → xs ++ ys ↭ xs ++ zs ++⁺ˡ [] ys↭zs = ys↭zs ++⁺ˡ (x ∷ xs) ys↭zs = ↭-prep x (++⁺ˡ xs ys↭zs) -++⁺ʳ : ∀ {xs ys : List A} zs → xs ↭ ys → xs ++ zs ↭ ys ++ zs +++⁺ʳ : ∀ zs → xs ↭ ys → xs ++ zs ↭ ys ++ zs ++⁺ʳ = Properties.++⁺ʳ ≈-refl ++⁺ : _++_ Preserves₂ _↭_ ⟶ _↭_ ⟶ _↭_ @@ -210,14 +235,14 @@ shift v≈w xs ys = Properties.shift ≈-refl ≈-sym ≈-trans v≈w xs {ys} -- Some other useful lemmas -zoom : ∀ h {t xs ys : List A} → xs ↭ ys → h ++ xs ++ t ↭ h ++ ys ++ t +zoom : ∀ h {t xs ys} → xs ↭ ys → h ++ xs ++ t ↭ h ++ ys ++ t zoom h {t} = ++⁺ˡ h ∘ ++⁺ʳ t -inject : ∀ (v : A) {ws xs ys zs} → ws ↭ ys → xs ↭ zs → +inject : ∀ v {ws xs ys zs} → ws ↭ ys → xs ↭ zs → ws ++ [ v ] ++ xs ↭ ys ++ [ v ] ++ zs inject v ws↭ys xs↭zs = ↭-trans (++⁺ˡ _ (↭-prep _ xs↭zs)) (++⁺ʳ _ ws↭ys) -shifts : ∀ xs ys {zs : List A} → xs ++ ys ++ zs ↭ ys ++ xs ++ zs +shifts : ∀ xs ys {zs} → xs ++ ys ++ zs ↭ ys ++ xs ++ zs shifts xs ys {zs} = begin xs ++ ys ++ zs ↭⟨ ++-assoc xs ys zs ⟨ (xs ++ ys) ++ zs ↭⟨ ++⁺ʳ zs (++-comm xs ys) ⟩ @@ -262,7 +287,7 @@ module _ {p} {P : Pred A p} (P? : Decidable P) (P≈ : P Respects _≈_) where module _ {p} {P : Pred A p} (P? : Decidable P) where - partition-↭ : ∀ xs → (let ys , zs = partition P? xs) → xs ↭ ys ++ zs + partition-↭ : ∀ xs → let ys , zs = partition P? xs in xs ↭ ys ++ zs partition-↭ [] = ↭-refl partition-↭ (x ∷ xs) with does (P? x) ... | true = ↭-prep x (partition-↭ xs) From c01559f34857737731fd6bc4f4bb013ca2ee11ea Mon Sep 17 00:00:00 2001 From: James McKinna Date: Mon, 11 Mar 2024 10:00:34 +0000 Subject: [PATCH 12/65] reverted changes --- src/Data/List/Relation/Binary/Pointwise/Properties.agda | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/Data/List/Relation/Binary/Pointwise/Properties.agda b/src/Data/List/Relation/Binary/Pointwise/Properties.agda index d203f915b2..1761dd3cda 100644 --- a/src/Data/List/Relation/Binary/Pointwise/Properties.agda +++ b/src/Data/List/Relation/Binary/Pointwise/Properties.agda @@ -9,7 +9,7 @@ module Data.List.Relation.Binary.Pointwise.Properties where open import Data.Product.Base using (_,_; uncurry) -open import Data.List.Base as List using (List; []; _∷_) +open import Data.List.Base using (List; []; _∷_) open import Level open import Relation.Binary.Core using (REL; _⇒_) open import Relation.Binary.Definitions @@ -25,8 +25,6 @@ private A : Set a B : Set b R S T : REL A B ℓ - xs : List A - ys : List B ------------------------------------------------------------------------ -- Relational properties From 747666f95221c5d539656910223eeb7634af952f Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 07:53:53 +0000 Subject: [PATCH 13/65] backport from `Propositional` --- .../Binary/Permutation/Setoid/Properties.agda | 65 +++++++++++-------- 1 file changed, 38 insertions(+), 27 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 86940bd619..4b0fdf642c 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -139,11 +139,8 @@ module _ (T : Setoid b ℓ) where map⁺ : ∀ {f} → f Preserves _≈_ ⟶ _≈′_ → xs ↭ ys → map f xs ↭′ map f ys map⁺ pres = Properties.map⁺ pres -{- - -- permutations preserve 'being a mapped list' - ↭-map-inv : map f xs ↭ zs → ∃ λ ys → zs ≡ map f ys × xs ↭ ys - ↭-map-inv p = --} + + ------------------------------------------------------------------------ -- _++_ @@ -266,6 +263,9 @@ dropMiddle : ∀ {vs} ws xs {ys zs} → dropMiddle {[]} ws xs p = p dropMiddle {v ∷ vs} ws xs p = dropMiddle ws xs (dropMiddleElement ws xs p) +drop-∷ : x ∷ xs ↭ x ∷ ys → xs ↭ ys +drop-∷ = dropMiddleElement [] [] + split-↭ : ∀ v as bs {xs} → xs ↭ as ++ [ v ] ++ bs → ∃₂ λ ps qs → xs ≋ ps ++ [ v ] ++ qs × ps ++ qs ↭ as ++ bs split-↭ v as bs p = Properties.split-↭ ≈-refl ≈-trans v as bs p @@ -293,30 +293,10 @@ module _ {p} {P : Pred A p} (P? : Decidable P) where ... | true = ↭-prep x (partition-↭ xs) ... | false = ↭-trans (↭-prep x (partition-↭ xs)) (↭-sym (↭-shift _ _)) - ------------------------------------------------------------------------- --- merge - -module _ {ℓ} {R : Rel A ℓ} (R? : B.Decidable R) where - - merge-↭ : ∀ xs ys → merge R? xs ys ↭ xs ++ ys - merge-↭ [] [] = ↭-refl - merge-↭ [] (y ∷ ys) = ↭-refl - merge-↭ (x ∷ xs) [] = ↭-sym (++-identityʳ (x ∷ xs)) - merge-↭ (x ∷ xs) (y ∷ ys) - with does (R? x y) | merge-↭ xs (y ∷ ys) | merge-↭ (x ∷ xs) ys - ... | true | rec | _ = ↭-prep x rec - ... | false | _ | rec = begin - y ∷ merge R? (x ∷ xs) ys <⟨ rec ⟩ - y ∷ x ∷ xs ++ ys ↭⟨ ↭-shift (x ∷ xs) ys ⟨ - (x ∷ xs) ++ y ∷ ys ≡⟨ List.++-assoc [ x ] xs (y ∷ ys) ⟨ - x ∷ xs ++ y ∷ ys ∎ - where open PermutationReasoning - ------------------------------------------------------------------------ -- _∷ʳ_ -∷↭∷ʳ : ∀ (x : A) xs → x ∷ xs ↭ xs ∷ʳ x +∷↭∷ʳ : ∀ x xs → x ∷ xs ↭ xs ∷ʳ x ∷↭∷ʳ x xs = ↭-sym (begin xs ++ [ x ] ↭⟨ ↭-shift xs [] ⟩ x ∷ xs ++ [] ≡⟨ List.++-identityʳ _ ⟩ @@ -326,10 +306,41 @@ module _ {ℓ} {R : Rel A ℓ} (R? : B.Decidable R) where ------------------------------------------------------------------------ -- ʳ++ -++↭ʳ++ : ∀ (xs ys : List A) → xs ++ ys ↭ xs ʳ++ ys +++↭ʳ++ : ∀ xs ys → xs ++ ys ↭ xs ʳ++ ys ++↭ʳ++ [] ys = ↭-refl ++↭ʳ++ (x ∷ xs) ys = ↭-trans (↭-sym (↭-shift xs ys)) (++↭ʳ++ xs (x ∷ ys)) +------------------------------------------------------------------------ +-- reverse + +↭-reverse : ∀ xs → reverse xs ↭ xs +↭-reverse [] = ↭-refl +↭-reverse (x ∷ xs) = begin + reverse (x ∷ xs) ≡⟨ List.unfold-reverse x xs ⟩ + reverse xs ∷ʳ x ↭⟨ ∷↭∷ʳ x (reverse xs) ⟨ + x ∷ reverse xs <⟨ ↭-reverse xs ⟩ + x ∷ xs ∎ + where open PermutationReasoning + +------------------------------------------------------------------------ +-- merge + +module _ {ℓ} {R : Rel A ℓ} (R? : B.Decidable R) where + + merge-↭ : ∀ xs ys → merge R? xs ys ↭ xs ++ ys + merge-↭ [] [] = ↭-refl + merge-↭ [] y∷ys@(_ ∷ _) = ↭-refl + merge-↭ x∷xs@(_ ∷ _) [] = ↭-sym (++-identityʳ x∷xs) + merge-↭ x∷xs@(x ∷ xs) y∷ys@(y ∷ ys) + with does (R? x y) | merge-↭ xs y∷ys | merge-↭ x∷xs ys + ... | true | rec | _ = ↭-prep x rec + ... | false | _ | rec = begin + y ∷ merge R? x∷xs ys <⟨ rec ⟩ + y ∷ x∷xs ++ ys ↭⟨ ↭-shift x∷xs ys ⟨ + x∷xs ++ y∷ys ≡⟨ List.++-assoc [ x ] xs y∷ys ⟨ + x∷xs ++ y∷ys ∎ + where open PermutationReasoning + ------------------------------------------------------------------------ -- foldr of Commutative Monoid From 985686ba1c0bbf88f67db3d88ab6accc323fb7c2 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 07:54:26 +0000 Subject: [PATCH 14/65] almost all the way there --- .../Permutation/Propositional/Properties.agda | 291 +++++------------- 1 file changed, 72 insertions(+), 219 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index f4a9eeacf1..f79587cfbd 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -1,7 +1,7 @@ ------------------------------------------------------------------------ -- The Agda standard library -- --- Properties of permutation +-- Properties of permutation in the Propositional case ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} @@ -39,114 +39,30 @@ private b p : Level B : Set b x y z : A - xs ys zs : List A + ws xs ys zs : List A + vs : List B open Propositional {A = A} -open Permutation (setoid A) public +open module P = Permutation (setoid A) public +-- legacy variations in naming + renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) +-- legacy variation in implicit/explicit parametrisation + hiding (shift) - ------------------------------------------------------------------------- --- Permutations of empty and singleton lists - -↭-singleton-inv : xs ↭ [ x ] → xs ≡ [ x ] -↭-singleton-inv ρ with _ , refl , refl ← ↭-singleton⁻¹ ρ = refl - ------------------------------------------------------------------------- --- sym -{- -↭-sym-involutive : ∀ {xs ys : List A} (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p -↭-sym-involutive refl = refl -↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭) -↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭) -↭-sym-involutive (trans ↭₁ ↭₂) = - cong₂ trans (↭-sym-involutive ↭₁) (↭-sym-involutive ↭₂) --} ------------------------------------------------------------------------ --- Relationships to other predicates -{- -Any-resp-[σ⁻¹∘σ] : {xs ys : List A} {P : Pred A p} → - (σ : xs ↭ ys) → - (ix : Any P xs) → - Any-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix -Any-resp-[σ⁻¹∘σ] refl ix = refl -Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ _) (there (here _)) = refl -Any-resp-[σ⁻¹∘σ] (trans σ₁ σ₂) ix - rewrite Any-resp-[σ⁻¹∘σ] σ₂ (Any-resp-↭ σ₁ ix) - rewrite Any-resp-[σ⁻¹∘σ] σ₁ ix - = refl -Any-resp-[σ⁻¹∘σ] (prep _ σ) (there ix) - rewrite Any-resp-[σ⁻¹∘σ] σ ix - = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) - rewrite Any-resp-[σ⁻¹∘σ] σ ix - = refl - -∈-resp-[σ⁻¹∘σ] : {xs ys : List A} {x : A} → - (σ : xs ↭ ys) → - (ix : x ∈ xs) → - ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix -∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] --} ------------------------------------------------------------------------- --- map -{- -module _ (f : A → B) where - - map⁺ : ∀ {xs ys} → xs ↭ ys → map f xs ↭ map f ys - map⁺ refl = refl - map⁺ (prep x p) = prep _ (map⁺ p) - map⁺ (swap x y p) = swap _ _ (map⁺ p) - map⁺ (trans p₁ p₂) = trans (map⁺ p₁) (map⁺ p₂) - - -- permutations preserve 'being a mapped list' - ↭-map-inv : ∀ {xs zs} → map f xs ↭ zs → ∃ λ ys → zs ≡ map f ys × xs ↭ ys - ↭-map-inv {[]} ρ = -, ↭-empty-inv (↭-sym ρ) , ↭-refl - ↭-map-inv {x ∷ []} ρ = -, ↭-singleton-inv (↭-sym ρ) , ↭-refl - ↭-map-inv {_ ∷ _ ∷ _} refl = -, refl , ↭-refl - ↭-map-inv {_ ∷ _ ∷ _} (prep _ ρ) with _ , refl , ρ′ ← ↭-map-inv ρ = -, refl , prep _ ρ′ - ↭-map-inv {_ ∷ _ ∷ _} (swap _ _ ρ) with _ , refl , ρ′ ← ↭-map-inv ρ = -, refl , swap _ _ ρ′ - ↭-map-inv {_ ∷ _ ∷ _} (trans ρ₁ ρ₂) with _ , refl , ρ₃ ← ↭-map-inv ρ₁ - with _ , refl , ρ₄ ← ↭-map-inv ρ₂ = -, refl , trans ρ₃ ρ₄ --} -{- ------------------------------------------------------------------------- --- _++_ - -++⁺ˡ : ∀ xs {ys zs : List A} → ys ↭ zs → xs ++ ys ↭ xs ++ zs -++⁺ˡ [] ys↭zs = ys↭zs -++⁺ˡ (x ∷ xs) ys↭zs = prep x (++⁺ˡ xs ys↭zs) - -++⁺ʳ : ∀ {xs ys : List A} zs → xs ↭ ys → xs ++ zs ↭ ys ++ zs -++⁺ʳ zs refl = refl -++⁺ʳ zs (prep x ↭) = prep x (++⁺ʳ zs ↭) -++⁺ʳ zs (swap x y ↭) = swap x y (++⁺ʳ zs ↭) -++⁺ʳ zs (trans ↭₁ ↭₂) = trans (++⁺ʳ zs ↭₁) (++⁺ʳ zs ↭₂) - -++⁺ : _++_ {A = A} Preserves₂ _↭_ ⟶ _↭_ ⟶ _↭_ -++⁺ ws↭xs ys↭zs = trans (++⁺ʳ _ ws↭xs) (++⁺ˡ _ ys↭zs) - -- Some useful lemmas -zoom : ∀ h {t xs ys : List A} → xs ↭ ys → h ++ xs ++ t ↭ h ++ ys ++ t -zoom h {t} = ++⁺ˡ h ∘ ++⁺ʳ t - -inject : ∀ (v : A) {ws xs ys zs} → ws ↭ ys → xs ↭ zs → - ws ++ [ v ] ++ xs ↭ ys ++ [ v ] ++ zs -inject v ws↭ys xs↭zs = trans (++⁺ˡ _ (prep v xs↭zs)) (++⁺ʳ _ ws↭ys) - shift : ∀ v (xs ys : List A) → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys -shift v [] ys = refl -shift v (x ∷ xs) ys = begin - x ∷ (xs ++ [ v ] ++ ys) <⟨ shift v xs ys ⟩ - x ∷ v ∷ xs ++ ys <<⟨ refl ⟩ - v ∷ x ∷ xs ++ ys ∎ +shift v = ↭-shift {v = v} + +{- not sure we need either of these for their proofs? +------------------------------------------------------------------------ +-- Some other useful lemmas, optimised for the Propositional case drop-mid-≡ : ∀ {x : A} ws xs {ys} {zs} → ws ++ [ x ] ++ ys ≡ xs ++ [ x ] ++ zs → ws ++ ys ↭ xs ++ zs -drop-mid-≡ [] [] eq with cong tail eq +drop-mid-≡ [] [] eq with cong List.tail eq drop-mid-≡ [] [] eq | refl = refl drop-mid-≡ [] (x ∷ xs) refl = shift _ xs _ drop-mid-≡ (w ∷ ws) [] refl = ↭-sym (shift _ ws _) @@ -194,140 +110,77 @@ drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ y ∷ xs) refl refl = swap y z (drop-mid′ p _ _ refl refl) drop-mid′ (trans p₁ p₂) ws xs refl refl with ∈-∃++ (∈-resp-↭ p₁ (∈-insert ws)) ... | (h , t , refl) = trans (drop-mid′ p₁ ws h refl refl) (drop-mid′ p₂ h xs refl refl) - --- Algebraic properties - -++-identityˡ : LeftIdentity {A = List A} _↭_ [] _++_ -++-identityˡ xs = refl - -++-identityʳ : RightIdentity {A = List A} _↭_ [] _++_ -++-identityʳ xs = ↭-reflexive (List.++-identityʳ xs) - -++-identity : Identity {A = List A} _↭_ [] _++_ -++-identity = ++-identityˡ , ++-identityʳ - -++-assoc : Associative {A = List A} _↭_ _++_ -++-assoc xs ys zs = ↭-reflexive (List.++-assoc xs ys zs) - -++-comm : Commutative {A = List A} _↭_ _++_ -++-comm [] ys = ↭-sym (++-identityʳ ys) -++-comm (x ∷ xs) ys = begin - x ∷ xs ++ ys <⟨ ++-comm xs ys ⟩ - x ∷ ys ++ xs ↭⟨ shift x ys xs ⟨ - ys ++ (x ∷ xs) ∎ - -++-isMagma : IsMagma {A = List A} _↭_ _++_ -++-isMagma = record - { isEquivalence = ↭-isEquivalence - ; ∙-cong = ++⁺ - } - -++-isSemigroup : IsSemigroup {A = List A} _↭_ _++_ -++-isSemigroup = record - { isMagma = ++-isMagma - ; assoc = ++-assoc - } - -++-isMonoid : IsMonoid {A = List A} _↭_ _++_ [] -++-isMonoid = record - { isSemigroup = ++-isSemigroup - ; identity = ++-identity - } - -++-isCommutativeMonoid : IsCommutativeMonoid {A = List A} _↭_ _++_ [] -++-isCommutativeMonoid = record - { isMonoid = ++-isMonoid - ; comm = ++-comm - } - -module _ {a} {A : Set a} where - - ++-magma : Magma _ _ - ++-magma = record - { isMagma = ++-isMagma {A = A} - } - - ++-semigroup : Semigroup a _ - ++-semigroup = record - { isSemigroup = ++-isSemigroup {A = A} - } - - ++-monoid : Monoid a _ - ++-monoid = record - { isMonoid = ++-isMonoid {A = A} - } - - ++-commutativeMonoid : CommutativeMonoid _ _ - ++-commutativeMonoid = record - { isCommutativeMonoid = ++-isCommutativeMonoid {A = A} - } - --- Another useful lemma - -shifts : ∀ xs ys {zs : List A} → xs ++ ys ++ zs ↭ ys ++ xs ++ zs -shifts xs ys {zs} = begin - xs ++ ys ++ zs ↭⟨ ++-assoc xs ys zs ⟨ - (xs ++ ys) ++ zs ↭⟨ ++⁺ʳ zs (++-comm xs ys) ⟩ - (ys ++ xs) ++ zs ↭⟨ ++-assoc ys xs zs ⟩ - ys ++ xs ++ zs ∎ +-} ------------------------------------------------------------------------ --- _∷_ - -drop-∷ : ∀ {x : A} {xs ys} → x ∷ xs ↭ x ∷ ys → xs ↭ ys -drop-∷ = drop-mid [] [] - +-- Additional/specialised properties which hold in the case _≈_ = _≡_ ------------------------------------------------------------------------ --- _∷ʳ_ - -∷↭∷ʳ : ∀ (x : A) xs → x ∷ xs ↭ xs ∷ʳ x -∷↭∷ʳ x xs = ↭-sym (begin - xs ++ [ x ] ↭⟨ shift x xs [] ⟩ - x ∷ xs ++ [] ≡⟨ List.++-identityʳ _ ⟩ - x ∷ xs ∎) ------------------------------------------------------------------------ --- ʳ++ +-- Permutations of singleton lists -++↭ʳ++ : ∀ (xs ys : List A) → xs ++ ys ↭ xs ʳ++ ys -++↭ʳ++ [] ys = ↭-refl -++↭ʳ++ (x ∷ xs) ys = ↭-trans (↭-sym (shift x xs ys)) (++↭ʳ++ xs (x ∷ ys)) +↭-singleton-inv : xs ↭ [ x ] → xs ≡ [ x ] +↭-singleton-inv ρ with _ , refl , refl ← ↭-singleton⁻¹ ρ = refl ------------------------------------------------------------------------ --- reverse - -↭-reverse : (xs : List A) → reverse xs ↭ xs -↭-reverse [] = ↭-refl -↭-reverse (x ∷ xs) = begin - reverse (x ∷ xs) ≡⟨ List.unfold-reverse x xs ⟩ - reverse xs ∷ʳ x ↭⟨ ∷↭∷ʳ x (reverse xs) ⟨ - x ∷ reverse xs <⟨ ↭-reverse xs ⟩ - x ∷ xs ∎ - where open PermutationReasoning - +-- sym +{- +↭-sym-involutive : ∀ {xs ys : List A} (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p +↭-sym-involutive refl = refl +↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭) +↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭) +↭-sym-involutive (trans ↭₁ ↭₂) = + cong₂ trans (↭-sym-involutive ↭₁) (↭-sym-involutive ↭₂) +-} ------------------------------------------------------------------------ --- merge +-- Relationships to other predicates +{- +Any-resp-[σ⁻¹∘σ] : {xs ys : List A} {P : Pred A p} → + (σ : xs ↭ ys) → + (ix : Any P xs) → + Any-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix +Any-resp-[σ⁻¹∘σ] refl ix = refl +Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl +Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl +Any-resp-[σ⁻¹∘σ] (swap _ _ _) (there (here _)) = refl +Any-resp-[σ⁻¹∘σ] (trans σ₁ σ₂) ix + rewrite Any-resp-[σ⁻¹∘σ] σ₂ (Any-resp-↭ σ₁ ix) + rewrite Any-resp-[σ⁻¹∘σ] σ₁ ix + = refl +Any-resp-[σ⁻¹∘σ] (prep _ σ) (there ix) + rewrite Any-resp-[σ⁻¹∘σ] σ ix + = refl +Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) + rewrite Any-resp-[σ⁻¹∘σ] σ ix + = refl -module _ {ℓ} {R : Rel A ℓ} (R? : Decidable R) where +∈-resp-[σ⁻¹∘σ] : {xs ys : List A} {x : A} → + (σ : xs ↭ ys) → + (ix : x ∈ xs) → + ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix +∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] +-} +------------------------------------------------------------------------ +-- map - merge-↭ : ∀ xs ys → merge R? xs ys ↭ xs ++ ys - merge-↭ [] [] = ↭-refl - merge-↭ [] (y ∷ ys) = ↭-refl - merge-↭ (x ∷ xs) [] = ↭-sym (++-identityʳ (x ∷ xs)) - merge-↭ (x ∷ xs) (y ∷ ys) - with does (R? x y) | merge-↭ xs (y ∷ ys) | merge-↭ (x ∷ xs) ys - ... | true | rec | _ = begin - x ∷ merge R? xs (y ∷ ys) <⟨ rec ⟩ - x ∷ xs ++ y ∷ ys ∎ - where open PermutationReasoning - ... | false | _ | rec = begin - y ∷ merge R? (x ∷ xs) ys <⟨ rec ⟩ - y ∷ x ∷ xs ++ ys ↭⟨ shift y (x ∷ xs) ys ⟨ - (x ∷ xs) ++ y ∷ ys ≡⟨ List.++-assoc [ x ] xs (y ∷ ys) ⟨ - x ∷ xs ++ y ∷ ys ∎ - where open PermutationReasoning +module _ {B : Set b} (f : A → B) where + open Propositional {A = B} using () renaming (_↭_ to _↭′_) + private module ↭′ = Permutation (setoid B) +{- + -- permutations preserve 'being a mapped list' + ↭-map-inv : List.map f xs ↭′ vs → ∃ λ ys → vs ≡ List.map f ys × xs ↭ ys + ↭-map-inv {xs = []} f*xs↭vs + with ≡.refl ← ↭′.↭-[]-invˡ f*xs↭vs = [] , ≡.refl , ↭-refl + ↭-map-inv {xs = x ∷ []} (Properties.refl f*xs≋vs) = {!f*xs≋vs!} + ↭-map-inv {xs = x ∷ []} (Properties.prep eq p) = {!!} + ↭-map-inv {xs = x ∷ []} (Properties.trans p p₁) = {!!} + ↭-map-inv {xs = x ∷ y ∷ xs} (Properties.refl x₁) = {!!} + ↭-map-inv {xs = x ∷ y ∷ xs} (Properties.prep eq p) = {!!} + ↭-map-inv {xs = x ∷ y ∷ xs} (Properties.swap eq₁ eq₂ p) = {!!} + ↭-map-inv {xs = x ∷ y ∷ xs} (Properties.trans p p₁) = {!!} -} + ------------------------------------------------------------------------ -- DEPRECATED NAMES ------------------------------------------------------------------------ From 865a962cafe527434e23759129233596a049556a Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 09:00:27 +0000 Subject: [PATCH 15/65] one(inductive) bug to fix; one cosmetic fix --- src/Data/List/Relation/Binary/BagAndSetEquality.agda | 6 +++--- .../Relation/Binary/Permutation/Setoid/Properties.agda | 9 ++++++--- 2 files changed, 9 insertions(+), 6 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 389721d4c5..4dc7a3c685 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -560,7 +560,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ------------------------------------------------------------------------ -- Relationships to other relations - +{- ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ ↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where @@ -584,10 +584,10 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = to∘from : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q to∘from p with from∘to (↭-sym p) ... | res rewrite ↭-sym-involutive p = res - +-} ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with empty-unique (↔-sym eq) -... | refl = refl +... | refl = ↭-refl ∼bag⇒↭ {A = A} {x ∷ xs} eq with ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) ... | zs₁ , zs₂ , p rewrite p = begin x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 4b0fdf642c..688b8564e4 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -263,9 +263,6 @@ dropMiddle : ∀ {vs} ws xs {ys zs} → dropMiddle {[]} ws xs p = p dropMiddle {v ∷ vs} ws xs p = dropMiddle ws xs (dropMiddleElement ws xs p) -drop-∷ : x ∷ xs ↭ x ∷ ys → xs ↭ ys -drop-∷ = dropMiddleElement [] [] - split-↭ : ∀ v as bs {xs} → xs ↭ as ++ [ v ] ++ bs → ∃₂ λ ps qs → xs ≋ ps ++ [ v ] ++ qs × ps ++ qs ↭ as ++ bs split-↭ v as bs p = Properties.split-↭ ≈-refl ≈-trans v as bs p @@ -274,6 +271,12 @@ split : ∀ (v : A) as bs {xs} → xs ↭ as ++ [ v ] ++ bs → ∃₂ λ ps qs split v as bs p with ps , qs , eq , _ ← split-↭ v as bs p = ps , qs , eq +------------------------------------------------------------------------ +-- _∷_ + +drop-∷ : x ∷ xs ↭ x ∷ ys → xs ↭ ys +drop-∷ = dropMiddleElement [] [] + ------------------------------------------------------------------------ -- filter From 159f6d602cc502dbbcca350811eb6f5ec481f412 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 09:18:28 +0000 Subject: [PATCH 16/65] two cosmetic knock-on fixes --- .../Relation/Ternary/Interleaving/Propositional.agda | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda index 55515cd447..87fb80e5e9 100644 --- a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda +++ b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda @@ -10,7 +10,7 @@ module Data.List.Relation.Ternary.Interleaving.Propositional {a} {A : Set a} whe open import Data.List.Base as List using (List; []; _∷_; _++_) open import Data.List.Relation.Binary.Permutation.Propositional as Perm using (_↭_) -open import Data.List.Relation.Binary.Permutation.Propositional.Properties using (shift) +open import Data.List.Relation.Binary.Permutation.Propositional.Properties using (↭-shift) import Data.List.Relation.Ternary.Interleaving.Setoid as General open import Relation.Binary.PropositionalEquality.Core using (refl) open import Relation.Binary.PropositionalEquality.Properties using (setoid) @@ -34,9 +34,9 @@ pattern consʳ xs = refl ∷ʳ xs -- New combinators toPermutation : ∀ {l r as} → Interleaving l r as → as ↭ l ++ r -toPermutation [] = Perm.refl -toPermutation (consˡ sp) = Perm.prep _ (toPermutation sp) +toPermutation [] = Perm.↭-refl +toPermutation (consˡ sp) = Perm.↭-prep _ (toPermutation sp) toPermutation {l} {r ∷ rs} {a ∷ as} (consʳ sp) = begin - a ∷ as ↭⟨ Perm.prep a (toPermutation sp) ⟩ - a ∷ l ++ rs ↭⟨ Perm.↭-sym (shift a l rs) ⟩ + a ∷ as <⟨ toPermutation sp ⟩ + a ∷ l ++ rs ↭⟨ ↭-shift l rs ⟨ l ++ a ∷ rs ∎ From 75a8414b9f659b4794c14713af658219872d7230 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 10:20:27 +0000 Subject: [PATCH 17/65] tidy up refactoring --- .../Binary/Permutation/Propositional.agda | 33 +++++++------------ 1 file changed, 11 insertions(+), 22 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 90b05d4efc..da092eb019 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -20,29 +20,18 @@ open import Relation.Binary.PropositionalEquality using (_≡_; refl; setoid) import Data.List.Relation.Binary.Permutation.Setoid as Permutation - ------------------------------------------------------------------------ --- An inductive definition of permutation - --- Note that one would expect that this would be defined in terms of --- `Permutation.Setoid`. This is not currently the case as it involves --- adding in a bunch of trivial `_≡_` proofs to the constructors which --- a) adds noise and b) prevents easy access to the variables `x`, `y`. --- This may be changed in future when a better solution is found. -{- -infix 3 _↭_ - -data _↭_ : Rel (List A) a where - refl : ∀ {xs} → xs ↭ xs - prep : ∀ {xs ys} x → xs ↭ ys → x ∷ xs ↭ x ∷ ys - swap : ∀ {xs ys} x y → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys - trans : ∀ {xs ys zs} → xs ↭ ys → ys ↭ zs → xs ↭ zs --} -open module ↭ = Permutation (setoid A) public - hiding ( ↭-reflexive; ↭-pointwise; ↭-trans - ; ↭-isEquivalence; ↭-setoid; module PermutationReasoning - ) +-- Definition of permutation + +private + module ↭ = Permutation (setoid A) +-- Note that this is now defined in terms of `Permutation.Setoid` (#2317) + +open ↭ public + hiding ( ↭-reflexive; ↭-pointwise; ↭-trans; ↭-isEquivalence + ; module PermutationReasoning; ↭-setoid + ) ------------------------------------------------------------------------ -- _↭_ is an equivalence @@ -74,7 +63,7 @@ open module ↭ = Permutation (setoid A) public -- A reasoning API to chain permutation proofs and allow "zooming in" -- to localised reasoning. For details of the axiomatisation, and -- in particular the refactored dependencies, --- see `Data.List.Relation.Binary.Permutation.Setoid` +-- see `Data.List.Relation.Binary.Permutation.Setoid.↭-Reasoning`. module PermutationReasoning = ↭-Reasoning ↭-isEquivalence ↭-pointwise From 89c9020fe5f3054da8b9e1563b16d60b2cde53ab Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 10:33:45 +0000 Subject: [PATCH 18/65] =?UTF-8?q?more=20level=20polymorphism=20in=20`map?= =?UTF-8?q?=E2=81=BA`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../List/Relation/Binary/Permutation/Setoid/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 688b8564e4..3aa5177264 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -131,7 +131,7 @@ steps-respʳ = Properties.steps-respʳ ≈-trans ------------------------------------------------------------------------ -- map -module _ (T : Setoid b ℓ) where +module _ (T : Setoid b r) where open Setoid T using () renaming (_≈_ to _≈′_) open Permutation T using () renaming (_↭_ to _↭′_) From 3e5230d3c73d50c64ca6cb5a5de874f1a5670f8f Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 10:37:53 +0000 Subject: [PATCH 19/65] =?UTF-8?q?fixed=20`map=E2=81=BA`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../Permutation/Propositional/Properties.agda | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index f79587cfbd..73426571b0 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -42,12 +42,15 @@ private ws xs ys zs : List A vs : List B -open Propositional {A = A} -open module P = Permutation (setoid A) public + module ↭ = Permutation (setoid A) + + +open Propositional {A = A} public +open ↭ public -- legacy variations in naming renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) -- legacy variation in implicit/explicit parametrisation - hiding (shift) + hiding (shift; map⁺) ------------------------------------------------------------------------ -- Some useful lemmas @@ -166,7 +169,9 @@ Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) module _ {B : Set b} (f : A → B) where open Propositional {A = B} using () renaming (_↭_ to _↭′_) - private module ↭′ = Permutation (setoid B) + + map⁺ : xs ↭ ys → List.map f xs ↭′ List.map f ys + map⁺ = ↭.map⁺ (setoid B) {!≡.cong f!} {- -- permutations preserve 'being a mapped list' ↭-map-inv : List.map f xs ↭′ vs → ∃ λ ys → vs ≡ List.map f ys × xs ↭ ys From 64c26b48d677c0d5858dd149a1725fa45a806bdb Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 10:58:21 +0000 Subject: [PATCH 20/65] updated `README`; need to refactor `PermutationReasoning` syntax --- .../List/Relation/Binary/Permutation.agda | 25 +++++++++++++------ 1 file changed, 18 insertions(+), 7 deletions(-) diff --git a/doc/README/Data/List/Relation/Binary/Permutation.agda b/doc/README/Data/List/Relation/Binary/Permutation.agda index e9141d89c7..a527249dd8 100644 --- a/doc/README/Data/List/Relation/Binary/Permutation.agda +++ b/doc/README/Data/List/Relation/Binary/Permutation.agda @@ -26,18 +26,18 @@ open import Relation.Binary.PropositionalEquality open import Data.List.Relation.Binary.Permutation.Propositional -- The permutation relation is written as `_↭_` and has four --- constructors. The first `refl` says that a list is always --- a permutation of itself, the second `prep` says that if the +-- *smart* constructors. The first `↭-refl` says that a list is +-- a permutation of itself, the second `↭-prep` says that if the -- heads of the lists are the same they can be skipped, the third --- `swap` says that the first two elements of the lists can be --- swapped and the fourth `trans` says that permutation proofs +-- `↭-swap` says that the first two elements of the lists can be +-- swapped and the fourth `↭trans` says that permutation proofs -- can be chained transitively. -- For example a proof that two lists are a permutation of one -- another can be written as follows: lem₁ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] -lem₁ = trans (prep 1 (swap 2 3 refl)) (swap 1 3 refl) +lem₁ = ↭-trans (↭-prep 1 (↭-swap 2 3 ↭-refl)) (↭-swap 1 3 ↭-refl) -- In practice it is difficult to parse the constructors in the -- proof above and hence understand why it holds. The @@ -48,10 +48,21 @@ open PermutationReasoning lem₂ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] lem₂ = begin - 1 ∷ 2 ∷ 3 ∷ [] ↭⟨ prep 1 (swap 2 3 refl) ⟩ - 1 ∷ 3 ∷ 2 ∷ [] ↭⟨ swap 1 3 refl ⟩ + 1 ∷ 2 ∷ 3 ∷ [] ↭⟨ ↭-prep 1 (↭-swap 2 3 ↭-refl) ⟩ + 1 ∷ 3 ∷ 2 ∷ [] ↭⟨ ↭-swap 1 3 ↭-refl ⟩ 3 ∷ 1 ∷ 2 ∷ [] ∎ +-- In practice it is further useful to extend `PermutationReasoning` +-- with specialised combinators `<⟨_⟩` and `<<⟨_⟩` to support the +-- distinguished use of the `↭-prep ` and `↭-prep ` steps, allowing +-- the above proof to be recast in the following form: +{- +lem₂ᵣ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] +lem₂ᵣ = begin + (1 ∷ 2 ∷ 3 ∷ []) <⟨ swap 2 3 ↭-refl ⟩ + (1 ∷ 3 ∷ 2 ∷ []) <<⟨ ↭-refl ⟩ + (3 ∷ 1 ∷ 2 ∷ []) ∎ +-} -- As might be expected, properties of the permutation relation may be -- found in: From 7642add43dc008a00e50680ad57a2398cde53307 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 11:12:22 +0000 Subject: [PATCH 21/65] BUG: `README`; need to refactor `PermutationReasoning` syntax? --- .../Data/List/Relation/Binary/Permutation.agda | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/doc/README/Data/List/Relation/Binary/Permutation.agda b/doc/README/Data/List/Relation/Binary/Permutation.agda index a527249dd8..3e03a983cf 100644 --- a/doc/README/Data/List/Relation/Binary/Permutation.agda +++ b/doc/README/Data/List/Relation/Binary/Permutation.agda @@ -54,15 +54,15 @@ lem₂ = begin -- In practice it is further useful to extend `PermutationReasoning` -- with specialised combinators `<⟨_⟩` and `<<⟨_⟩` to support the --- distinguished use of the `↭-prep ` and `↭-prep ` steps, allowing +-- distinguished use of the `↭-prep ` and `↭-swap` steps, allowing -- the above proof to be recast in the following form: -{- + lem₂ᵣ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] lem₂ᵣ = begin - (1 ∷ 2 ∷ 3 ∷ []) <⟨ swap 2 3 ↭-refl ⟩ - (1 ∷ 3 ∷ 2 ∷ []) <<⟨ ↭-refl ⟩ - (3 ∷ 1 ∷ 2 ∷ []) ∎ --} + 1 ∷ 2 ∷ 3 ∷ [] <⟨ swap 2 3 ↭-refl ⟩ + 1 ∷ 3 ∷ 2 ∷ [] <<⟨ ↭-refl ⟩ + 3 ∷ 1 ∷ 2 ∷ [] ∎ + -- As might be expected, properties of the permutation relation may be -- found in: From cbbda53498a18a94124335a0ed838cadf214f58d Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 11:19:56 +0000 Subject: [PATCH 22/65] uncommitted changes --- .../Relation/Binary/Permutation/Propositional/Properties.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 73426571b0..9fba138b31 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -47,7 +47,7 @@ private open Propositional {A = A} public open ↭ public --- legacy variations in naming +-- legacy variation in naming renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) -- legacy variation in implicit/explicit parametrisation hiding (shift; map⁺) @@ -171,7 +171,7 @@ module _ {B : Set b} (f : A → B) where open Propositional {A = B} using () renaming (_↭_ to _↭′_) map⁺ : xs ↭ ys → List.map f xs ↭′ List.map f ys - map⁺ = ↭.map⁺ (setoid B) {!≡.cong f!} + map⁺ = ↭.map⁺ (setoid B) (≡.cong f) {- -- permutations preserve 'being a mapped list' ↭-map-inv : List.map f xs ↭′ vs → ∃ λ ys → vs ≡ List.map f ys × xs ↭ ys From 55b4ff5c327ae125dd6fc4f10117083161c4c128 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 11:42:15 +0000 Subject: [PATCH 23/65] improved use of `variable`s --- .../Binary/Permutation/Setoid/Properties.agda | 35 ++++++++++--------- 1 file changed, 19 insertions(+), 16 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 3aa5177264..0d67d3e332 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -17,9 +17,6 @@ open import Data.Bool.Base using (true; false) open import Data.List.Base open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise) -import Data.List.Relation.Binary.Equality.Setoid as Equality -import Data.List.Relation.Binary.Permutation.Setoid as Permutation -import Data.List.Relation.Binary.Permutation.Homogeneous as Properties open import Data.List.Relation.Unary.Any as Any using (Any; here; there) open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) @@ -41,6 +38,10 @@ open import Relation.Nullary.Decidable using (yes; no; does) open import Relation.Nullary.Negation using (¬_; contradiction) open import Relation.Unary using (Pred; Decidable) +import Data.List.Relation.Binary.Equality.Setoid as Equality +import Data.List.Relation.Binary.Permutation.Setoid as Permutation +import Data.List.Relation.Binary.Permutation.Homogeneous as Properties + open Setoid S using (_≈_) renaming (Carrier to A; refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) @@ -55,21 +56,23 @@ private b p r : Level xs ys zs : List A x y z v w : A + P : Pred A p + R : Rel A r ------------------------------------------------------------------------ -- Relationships to other predicates ------------------------------------------------------------------------ -All-resp-↭ : ∀ {P : Pred A p} → P Respects _≈_ → (All P) Respects _↭_ +All-resp-↭ : P Respects _≈_ → (All P) Respects _↭_ All-resp-↭ = Properties.All-resp-↭ -Any-resp-↭ : ∀ {P : Pred A p} → P Respects _≈_ → (Any P) Respects _↭_ +Any-resp-↭ : P Respects _≈_ → (Any P) Respects _↭_ Any-resp-↭ = Properties.Any-resp-↭ -AllPairs-resp-↭ : ∀ {R : Rel A r} → Symmetric R → R Respects₂ _≈_ → (AllPairs R) Respects _↭_ +AllPairs-resp-↭ : Symmetric R → R Respects₂ _≈_ → (AllPairs R) Respects _↭_ AllPairs-resp-↭ = Properties.AllPairs-resp-↭ -∈-resp-↭ : ∀ {x} → (x ∈_) Respects _↭_ +∈-resp-↭ : (x ∈_) Respects _↭_ ∈-resp-↭ = Any-resp-↭ (flip ≈-trans) Unique-resp-↭ : Unique Respects _↭_ @@ -92,14 +95,14 @@ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ -- Properties of steps ------------------------------------------------------------------------ -0 Date: Tue, 12 Mar 2024 12:01:20 +0000 Subject: [PATCH 24/65] fixed `import`s to reflect `public` export of the relation in `Propositional.Properties` --- src/Data/List/Relation/Binary/BagAndSetEquality.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 4dc7a3c685..256b982570 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -22,7 +22,6 @@ open import Data.List.Membership.Propositional using (_∈_) open import Data.List.Membership.Propositional.Properties open import Data.List.Relation.Binary.Subset.Propositional.Properties using (⊆-preorder) -open import Data.List.Relation.Binary.Permutation.Propositional open import Data.List.Relation.Binary.Permutation.Propositional.Properties open import Data.Product.Base as Prod hiding (map) import Data.Product.Function.Dependent.Propositional as Σ From 4e2fb04a7704fd806f85f4a831cedc919efc9520 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 12:37:32 +0000 Subject: [PATCH 25/65] =?UTF-8?q?tightened=20the=20paramterisation=20of=20?= =?UTF-8?q?`=E2=86=AD-shift`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../Relation/Binary/BagAndSetEquality.agda | 25 ++++++++----------- .../Permutation/Propositional/Properties.agda | 16 ++++++------ .../Binary/Permutation/Setoid/Properties.agda | 14 +++++------ 3 files changed, 27 insertions(+), 28 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 256b982570..578442c0bd 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -537,21 +537,18 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = lemma : ∀ {xs ys} (inv : x ∷ xs ∼[ bag ] x ∷ ys) {z} → Well-behaved (Inverse.to (∼→⊎↔⊎ inv {z})) - lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ with - zero ≡⟨⟩ - index-of {xs = x ∷ xs} (here p) ≡⟨⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ ≡.sym $ - to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟩ + lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ with begin + zero ≡⟨⟩ + index-of {xs = x ∷ xs} (here p) ≡⟨⟩ + index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ ≡.sym $ to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟩ index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩ - index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩ - index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ ≡.cong index-of $ ≡.sym $ - strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ - to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) ≡⟨⟩ - index-of {xs = x ∷ xs} (there z∈xs) ≡⟨⟩ - suc (index-of {xs = xs} z∈xs) ∎ + index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩ + index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ ≡.cong index-of $ ≡.sym $ strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟩ + index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩ + index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) ≡⟨⟩ + index-of {xs = x ∷ xs} (there z∈xs) ≡⟨⟩ + suc (index-of {xs = xs} z∈xs) ∎ where open Inverse open ≡-Reasoning @@ -591,7 +588,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ... | zs₁ , zs₂ , p rewrite p = begin x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ - x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨ + x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-shift zs₁ zs₂ ⟨ zs₁ ++ x ∷ zs₂ ∎ where open CommutativeMonoid (commutativeMonoid bag A) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 9fba138b31..60e0f0a099 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -52,15 +52,9 @@ open ↭ public -- legacy variation in implicit/explicit parametrisation hiding (shift; map⁺) ------------------------------------------------------------------------- --- Some useful lemmas - -shift : ∀ v (xs ys : List A) → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys -shift v = ↭-shift {v = v} - {- not sure we need either of these for their proofs? ------------------------------------------------------------------------ --- Some other useful lemmas, optimised for the Propositional case +-- Some useful lemmas, optimised for the Propositional case drop-mid-≡ : ∀ {x : A} ws xs {ys} {zs} → ws ++ [ x ] ++ ys ≡ xs ++ [ x ] ++ zs → @@ -200,4 +194,12 @@ module _ {B : Set b} (f : A → B) where Please use ↭-[]-invʳ instead." #-} +shift : ∀ v xs ys → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys +shift v xs ys = ↭-shift {v = v} xs {ys} +{-# WARNING_ON_USAGE shift +"Warning: shift was deprecated in v2.1. +Please use ↭-shift instead. \\ +NB. Parametrisation has changed to make v and ys implicit." +#-} + diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 0d67d3e332..76c74d69d0 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -150,8 +150,8 @@ module _ (T : Setoid b r) where shift : v ≈ w → ∀ xs ys → xs ++ [ v ] ++ ys ↭ w ∷ xs ++ ys shift v≈w xs ys = Properties.shift ≈-refl ≈-sym ≈-trans v≈w xs {ys} -↭-shift : ∀ xs ys → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys -↭-shift = shift ≈-refl +↭-shift : ∀ xs {ys} → xs ++ [ v ] ++ ys ↭ v ∷ xs ++ ys +↭-shift xs {ys} = shift ≈-refl xs ys ++⁺ˡ : ∀ xs {ys zs} → ys ↭ zs → xs ++ ys ↭ xs ++ zs ++⁺ˡ [] ys↭zs = ys↭zs @@ -181,7 +181,7 @@ shift v≈w xs ys = Properties.shift ≈-refl ≈-sym ≈-trans v≈w xs {ys} ++-comm [] ys = ↭-sym (++-identityʳ ys) ++-comm (x ∷ xs) ys = begin x ∷ xs ++ ys <⟨ ++-comm xs ys ⟩ - x ∷ ys ++ xs ↭⟨ ↭-shift ys xs ⟨ + x ∷ ys ++ xs ↭⟨ ↭-shift ys ⟨ ys ++ (x ∷ xs) ∎ where open PermutationReasoning @@ -297,14 +297,14 @@ module _ (P? : Decidable P) where partition-↭ [] = ↭-refl partition-↭ (x ∷ xs) with does (P? x) ... | true = ↭-prep x (partition-↭ xs) - ... | false = ↭-trans (↭-prep x (partition-↭ xs)) (↭-sym (↭-shift _ _)) + ... | false = ↭-trans (↭-prep x (partition-↭ xs)) (↭-sym (↭-shift _)) ------------------------------------------------------------------------ -- _∷ʳ_ ∷↭∷ʳ : ∀ x xs → x ∷ xs ↭ xs ∷ʳ x ∷↭∷ʳ x xs = ↭-sym (begin - xs ++ [ x ] ↭⟨ ↭-shift xs [] ⟩ + xs ++ [ x ] ↭⟨ ↭-shift xs ⟩ x ∷ xs ++ [] ≡⟨ List.++-identityʳ _ ⟩ x ∷ xs ∎) where open PermutationReasoning @@ -314,7 +314,7 @@ module _ (P? : Decidable P) where ++↭ʳ++ : ∀ xs ys → xs ++ ys ↭ xs ʳ++ ys ++↭ʳ++ [] ys = ↭-refl -++↭ʳ++ (x ∷ xs) ys = ↭-trans (↭-sym (↭-shift xs ys)) (++↭ʳ++ xs (x ∷ ys)) +++↭ʳ++ (x ∷ xs) ys = ↭-trans (↭-sym (↭-shift xs)) (++↭ʳ++ xs (x ∷ ys)) ------------------------------------------------------------------------ -- reverse @@ -342,7 +342,7 @@ module _ (R? : B.Decidable R) where ... | true | rec | _ = ↭-prep x rec ... | false | _ | rec = begin y ∷ merge R? x∷xs ys <⟨ rec ⟩ - y ∷ x∷xs ++ ys ↭⟨ ↭-shift x∷xs ys ⟨ + y ∷ x∷xs ++ ys ↭⟨ ↭-shift x∷xs ⟨ x∷xs ++ y∷ys ≡⟨ List.++-assoc [ x ] xs y∷ys ⟨ x∷xs ++ y∷ys ∎ where open PermutationReasoning From e652a80a284b45a01b8643ffad0dff69b2844697 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 12:39:11 +0000 Subject: [PATCH 26/65] knock-on viscosity --- src/Data/List/Relation/Binary/BagAndSetEquality.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 578442c0bd..2650ba794a 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -588,7 +588,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ... | zs₁ , zs₂ , p rewrite p = begin x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ - x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-shift zs₁ zs₂ ⟨ + x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-shift zs₁ ⟨ zs₁ ++ x ∷ zs₂ ∎ where open CommutativeMonoid (commutativeMonoid bag A) From 542caf4a1ab5fa85e004bba26897d3138519c4a9 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 13:04:18 +0000 Subject: [PATCH 27/65] knock-on viscosity --- src/Data/List/Relation/Ternary/Interleaving/Propositional.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda index 87fb80e5e9..60cec992fd 100644 --- a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda +++ b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda @@ -38,5 +38,5 @@ toPermutation [] = Perm.↭-refl toPermutation (consˡ sp) = Perm.↭-prep _ (toPermutation sp) toPermutation {l} {r ∷ rs} {a ∷ as} (consʳ sp) = begin a ∷ as <⟨ toPermutation sp ⟩ - a ∷ l ++ rs ↭⟨ ↭-shift l rs ⟨ + a ∷ l ++ rs ↭⟨ ↭-shift l ⟨ l ++ a ∷ rs ∎ From 8bda5d6284b5c087317ad8a2fb4cb3554682f92a Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 14:31:03 +0000 Subject: [PATCH 28/65] commented out bug in reasoning --- doc/README/Data/List/Relation/Binary/Permutation.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/doc/README/Data/List/Relation/Binary/Permutation.agda b/doc/README/Data/List/Relation/Binary/Permutation.agda index 3e03a983cf..eed81a3645 100644 --- a/doc/README/Data/List/Relation/Binary/Permutation.agda +++ b/doc/README/Data/List/Relation/Binary/Permutation.agda @@ -56,13 +56,13 @@ lem₂ = begin -- with specialised combinators `<⟨_⟩` and `<<⟨_⟩` to support the -- distinguished use of the `↭-prep ` and `↭-swap` steps, allowing -- the above proof to be recast in the following form: - +{- lem₂ᵣ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] lem₂ᵣ = begin 1 ∷ 2 ∷ 3 ∷ [] <⟨ swap 2 3 ↭-refl ⟩ 1 ∷ 3 ∷ 2 ∷ [] <<⟨ ↭-refl ⟩ 3 ∷ 1 ∷ 2 ∷ [] ∎ - +-} -- As might be expected, properties of the permutation relation may be -- found in: From 331220d0043251767d25c7a515697b4840cb847a Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 14:31:53 +0000 Subject: [PATCH 29/65] more properties needing specialised --- .../Permutation/Propositional/Properties.agda | 27 +++++++++++++++---- 1 file changed, 22 insertions(+), 5 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 60e0f0a099..d56e30e3c6 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -17,6 +17,8 @@ open import Data.Product.Base using (-,_; proj₂) open import Data.List.Base as List using (List; []; _∷_; [_]; _++_) open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.All using (All; []; _∷_) +open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +import Data.List.Relation.Unary.Unique.Setoid as Unique open import Data.List.Membership.Propositional open import Data.List.Membership.Propositional.Properties import Data.List.Properties as List @@ -32,15 +34,16 @@ open import Relation.Nullary import Data.List.Relation.Binary.Permutation.Propositional as Propositional import Data.List.Relation.Binary.Permutation.Setoid.Properties as Permutation -import Data.List.Relation.Binary.Permutation.Homogeneous as Properties private variable - b p : Level + b p r : Level B : Set b x y z : A ws xs ys zs : List A vs : List B + P : Pred A p + R : Rel A r module ↭ = Permutation (setoid A) @@ -50,9 +53,11 @@ open ↭ public -- legacy variation in naming renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) -- legacy variation in implicit/explicit parametrisation - hiding (shift; map⁺) + hiding (shift) +-- needing to specialise to ≡, where `Respects` and `Preserves` are trivial + hiding (map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭) -{- not sure we need either of these for their proofs? +{- not sure we need either of these for their proofs? so renaming above ------------------------------------------------------------------------ -- Some useful lemmas, optimised for the Propositional case @@ -131,11 +136,23 @@ drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl -} ------------------------------------------------------------------------ -- Relationships to other predicates +------------------------------------------------------------------------ + +All-resp-↭ : (All P) Respects _↭_ +All-resp-↭ = ↭.All-resp-↭ (≡.resp _) + +Any-resp-↭ : (Any P) Respects _↭_ +Any-resp-↭ = ↭.Any-resp-↭ (≡.resp _) + +∈-resp-↭ : (x ∈_) Respects _↭_ +∈-resp-↭ = Any-resp-↭ + {- Any-resp-[σ⁻¹∘σ] : {xs ys : List A} {P : Pred A p} → (σ : xs ↭ ys) → (ix : Any P xs) → - Any-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix + Any-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix +Any-resp-[σ⁻¹∘σ] = {!!} Any-resp-[σ⁻¹∘σ] refl ix = refl Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl From 79e0c8f538a802d72c78e48bbb63feb3651e751d Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 16:11:46 +0000 Subject: [PATCH 30/65] allow unsolved metas; nearly there! --- .../Relation/Binary/BagAndSetEquality.agda | 40 ++--- .../Permutation/Propositional/Properties.agda | 148 +++++++++--------- .../Subset/Propositional/Properties.agda | 2 +- 3 files changed, 95 insertions(+), 95 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 2650ba794a..9d4a50499e 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -4,7 +4,7 @@ -- Bag and set equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible #-} module Data.List.Relation.Binary.BagAndSetEquality where @@ -556,17 +556,19 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ------------------------------------------------------------------------ -- Relationships to other relations -{- + ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ ↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where - to : ∀ {xs ys} → xs ↭ ys → v ∈ xs → v ∈ ys - to xs↭ys = Any-resp-↭ {A = A} xs↭ys + to : xs ↭ ys → v ∈ xs → v ∈ ys + to = ∈-resp-↭ - from : ∀ {xs ys} → xs ↭ ys → v ∈ ys → v ∈ xs - from xs↭ys = Any-resp-↭ (↭-sym xs↭ys) + from : xs ↭ ys → v ∈ ys → v ∈ xs + from = ∈-resp-↭ ∘ ↭-sym - from∘to : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ xs) → from p (to p q) ≡ q + from∘to : (p : xs ↭ ys) (q : v ∈ xs) → from p (to p q) ≡ q + from∘to p v∈xs = {!p!} +{- from∘to refl v∈xs = refl from∘to (prep _ _) (here refl) = refl from∘to (prep _ p) (there v∈xs) = ≡.cong there (from∘to p v∈xs) @@ -574,22 +576,20 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = from∘to (swap x y p) (there (here refl)) = refl from∘to (swap x y p) (there (there v∈xs)) = ≡.cong (there ∘ there) (from∘to p v∈xs) from∘to (trans p₁ p₂) v∈xs - rewrite from∘to p₂ (Any-resp-↭ p₁ v∈xs) + rewrite from∘to p₂ (∈-resp-↭ p₁ v∈xs) | from∘to p₁ v∈xs = refl - - to∘from : ∀ {xs ys} (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q - to∘from p with from∘to (↭-sym p) - ... | res rewrite ↭-sym-involutive p = res -} + to∘from : (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q + to∘from p with res ← from∘to (↭-sym p) rewrite ↭-sym-involutive p = res + ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} -∼bag⇒↭ {A = A} {[]} eq with empty-unique (↔-sym eq) -... | refl = ↭-refl -∼bag⇒↭ {A = A} {x ∷ xs} eq with ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) -... | zs₁ , zs₂ , p rewrite p = begin - x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ - x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ - x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-shift zs₁ ⟨ - zs₁ ++ x ∷ zs₂ ∎ +∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl +∼bag⇒↭ {A = A} {x ∷ xs} eq + with zs₁ , zs₂ , refl ← ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) = begin + x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩ + x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩ + x ∷ (zs₁ ++ zs₂) ↭⟨ ↭-shift zs₁ ⟨ + zs₁ ++ x ∷ zs₂ ∎ where open CommutativeMonoid (commutativeMonoid bag A) open PermutationReasoning diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index d56e30e3c6..4854064665 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -4,7 +4,7 @@ -- Properties of permutation in the Propositional case ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible --allow-unsolved-metas #-} module Data.List.Relation.Binary.Permutation.Propositional.Properties {a} {A : Set a} where @@ -48,72 +48,19 @@ private module ↭ = Permutation (setoid A) +------------------------------------------------------------------------ +-- Export all Setoid lemmas which hold unchanged in the case _≈_ = _≡_ +------------------------------------------------------------------------ + open Propositional {A = A} public open ↭ public --- legacy variation in naming +-- legacy variations in naming renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) --- legacy variation in implicit/explicit parametrisation +-- DEPRECATION: legacy variation in implicit/explicit parametrisation hiding (shift) -- needing to specialise to ≡, where `Respects` and `Preserves` are trivial hiding (map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭) -{- not sure we need either of these for their proofs? so renaming above ------------------------------------------------------------------------- --- Some useful lemmas, optimised for the Propositional case - -drop-mid-≡ : ∀ {x : A} ws xs {ys} {zs} → - ws ++ [ x ] ++ ys ≡ xs ++ [ x ] ++ zs → - ws ++ ys ↭ xs ++ zs -drop-mid-≡ [] [] eq with cong List.tail eq -drop-mid-≡ [] [] eq | refl = refl -drop-mid-≡ [] (x ∷ xs) refl = shift _ xs _ -drop-mid-≡ (w ∷ ws) [] refl = ↭-sym (shift _ ws _) -drop-mid-≡ (w ∷ ws) (x ∷ xs) eq with List.∷-injective eq -... | refl , eq′ = prep w (drop-mid-≡ ws xs eq′) - -drop-mid : ∀ {x : A} ws xs {ys zs} → - ws ++ [ x ] ++ ys ↭ xs ++ [ x ] ++ zs → - ws ++ ys ↭ xs ++ zs -drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl - where - drop-mid′ : ∀ {l′ l″ : List A} → l′ ↭ l″ → - ∀ ws xs {ys zs} → - ws ++ [ x ] ++ ys ≡ l′ → - xs ++ [ x ] ++ zs ≡ l″ → - ws ++ ys ↭ xs ++ zs - drop-mid′ refl ws xs refl eq = drop-mid-≡ ws xs (≡.sym eq) - drop-mid′ (prep x p) [] [] refl eq with cong tail eq - drop-mid′ (prep x p) [] [] refl eq | refl = p - drop-mid′ (prep x p) [] (x ∷ xs) refl refl = trans p (shift _ _ _) - drop-mid′ (prep x p) (w ∷ ws) [] refl refl = trans (↭-sym (shift _ _ _)) p - drop-mid′ (prep x p) (w ∷ ws) (x ∷ xs) refl refl = prep _ (drop-mid′ p ws xs refl refl) - drop-mid′ (swap y z p) [] [] refl refl = prep _ p - drop-mid′ (swap y z p) [] (x ∷ []) refl eq with cong {B = List _} - (λ { (x ∷ _ ∷ xs) → x ∷ xs - ; _ → [] - }) - eq - drop-mid′ (swap y z p) [] (x ∷ []) refl eq | refl = prep _ p - drop-mid′ (swap y z p) [] (x ∷ _ ∷ xs) refl refl = prep _ (trans p (shift _ _ _)) - drop-mid′ (swap y z p) (w ∷ []) [] refl eq with cong tail eq - drop-mid′ (swap y z p) (w ∷ []) [] refl eq | refl = prep _ p - drop-mid′ (swap y z p) (w ∷ x ∷ ws) [] refl refl = prep _ (trans (↭-sym (shift _ _ _)) p) - drop-mid′ (swap y y p) (y ∷ []) (y ∷ []) refl refl = prep _ p - drop-mid′ (swap y z p) (y ∷ []) (z ∷ y ∷ xs) refl refl = begin - _ ∷ _ <⟨ p ⟩ - _ ∷ (xs ++ _ ∷ _) <⟨ shift _ _ _ ⟩ - _ ∷ _ ∷ xs ++ _ <<⟨ refl ⟩ - _ ∷ _ ∷ xs ++ _ ∎ - drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ []) refl refl = begin - _ ∷ _ ∷ ws ++ _ <<⟨ refl ⟩ - _ ∷ (_ ∷ ws ++ _) <⟨ ↭-sym (shift _ _ _) ⟩ - _ ∷ (ws ++ _ ∷ _) <⟨ p ⟩ - _ ∷ _ ∎ - drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ y ∷ xs) refl refl = swap y z (drop-mid′ p _ _ refl refl) - drop-mid′ (trans p₁ p₂) ws xs refl refl with ∈-∃++ (∈-resp-↭ p₁ (∈-insert ws)) - ... | (h , t , refl) = trans (drop-mid′ p₁ ws h refl refl) (drop-mid′ p₂ h xs refl refl) --} - ------------------------------------------------------------------------ -- Additional/specialised properties which hold in the case _≈_ = _≡_ ------------------------------------------------------------------------ @@ -126,8 +73,10 @@ drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl ------------------------------------------------------------------------ -- sym + +↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p +↭-sym-involutive p = {!p!} {- -↭-sym-involutive : ∀ {xs ys : List A} (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p ↭-sym-involutive refl = refl ↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭) ↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭) @@ -147,12 +96,10 @@ Any-resp-↭ = ↭.Any-resp-↭ (≡.resp _) ∈-resp-↭ : (x ∈_) Respects _↭_ ∈-resp-↭ = Any-resp-↭ -{- -Any-resp-[σ⁻¹∘σ] : {xs ys : List A} {P : Pred A p} → - (σ : xs ↭ ys) → - (ix : Any P xs) → +Any-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : Any P xs) → Any-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix -Any-resp-[σ⁻¹∘σ] = {!!} +Any-resp-[σ⁻¹∘σ] p ix = {!p!} +{- Any-resp-[σ⁻¹∘σ] refl ix = refl Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl @@ -167,13 +114,11 @@ Any-resp-[σ⁻¹∘σ] (prep _ σ) (there ix) Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) rewrite Any-resp-[σ⁻¹∘σ] σ ix = refl - -∈-resp-[σ⁻¹∘σ] : {xs ys : List A} {x : A} → - (σ : xs ↭ ys) → - (ix : x ∈ xs) → - ∈-resp-↭ (trans σ (↭-sym σ)) ix ≡ ix -∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] -} +∈-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : x ∈ xs) → + ∈-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix +∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] + ------------------------------------------------------------------------ -- map @@ -197,6 +142,63 @@ module _ {B : Set b} (f : A → B) where ↭-map-inv {xs = x ∷ y ∷ xs} (Properties.trans p p₁) = {!!} -} +------------------------------------------------------------------------ +-- Some other useful lemmas, optimised for the Propositional case +{- not sure we need these for their proofs? so renamed on import above + +drop-mid-≡ : ∀ {v : A} ws xs {ys} {zs} → + ws ++ [ x ] ++ ys ≡ xs ++ [ v ] ++ zs → + ws ++ ys ↭ xs ++ zs +drop-mid-≡ [] [] ys≡zs + with refl ← cong List.tail ys≡zs = refl +drop-mid-≡ [] (x ∷ xs) refl = ↭-shift xs +drop-mid-≡ (w ∷ ws) [] refl = ↭-sym (↭-shift ws) +drop-mid-≡ (w ∷ ws) (x ∷ xs) w∷ws[v]ys≡x∷xs[v]zs + with refl , ws[v]ys≡xs[v]zs ← List.∷-injective eq + = prep w (drop-mid-≡ ws xs ws[v]ys≡xs[v]zs) + +drop-mid : ∀ {v : A} ws xs {ys zs} → + ws ++ [ v ] ++ ys ↭ xs ++ [ v ] ++ zs → + ws ++ ys ↭ xs ++ zs +drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl + where + drop-mid′ : ∀ {l′ l″ : List A} → l′ ↭ l″ → + ∀ ws xs {ys zs} → + ws ++ [ x ] ++ ys ≡ l′ → + xs ++ [ x ] ++ zs ≡ l″ → + ws ++ ys ↭ xs ++ zs + drop-mid′ refl ws xs refl eq = drop-mid-≡ ws xs (≡.sym eq) + drop-mid′ (prep x p) [] [] refl eq with refl ← cong tail eq = p + drop-mid′ (prep x p) [] (x ∷ xs) refl refl = trans p (shift _ _ _) + drop-mid′ (prep x p) (w ∷ ws) [] refl refl = trans (↭-sym (shift _ _ _)) p + drop-mid′ (prep x p) (w ∷ ws) (x ∷ xs) refl refl = prep _ (drop-mid′ p ws xs refl refl) + drop-mid′ (swap y z p) [] [] refl refl = prep _ p + drop-mid′ (swap y z p) [] (x ∷ []) refl eq with cong {B = List _} + (λ { (x ∷ _ ∷ xs) → x ∷ xs + ; _ → [] + }) + eq + drop-mid′ (swap y z p) [] (x ∷ []) refl eq | refl = prep _ p + drop-mid′ (swap y z p) [] (x ∷ _ ∷ xs) refl refl = prep _ (trans p (shift _ _ _)) + drop-mid′ (swap y z p) (w ∷ []) [] refl eq with cong tail eq + drop-mid′ (swap y z p) (w ∷ []) [] refl eq | refl = prep _ p + drop-mid′ (swap y z p) (w ∷ x ∷ ws) [] refl refl = prep _ (trans (↭-sym (shift _ _ _)) p) + drop-mid′ (swap y y p) (y ∷ []) (y ∷ []) refl refl = prep _ p + drop-mid′ (swap y z p) (y ∷ []) (z ∷ y ∷ xs) refl refl = begin + _ ∷ _ <⟨ p ⟩ + _ ∷ (xs ++ _ ∷ _) <⟨ shift _ _ _ ⟩ + _ ∷ _ ∷ xs ++ _ <<⟨ refl ⟩ + _ ∷ _ ∷ xs ++ _ ∎ + drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ []) refl refl = begin + _ ∷ _ ∷ ws ++ _ <<⟨ refl ⟩ + _ ∷ (_ ∷ ws ++ _) <⟨ ↭-sym (shift _ _ _) ⟩ + _ ∷ (ws ++ _ ∷ _) <⟨ p ⟩ + _ ∷ _ ∎ + drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ y ∷ xs) refl refl = swap y z (drop-mid′ p _ _ refl refl) + drop-mid′ (trans p₁ p₂) ws xs refl refl with ∈-∃++ (∈-resp-↭ p₁ (∈-insert ws)) + ... | (h , t , refl) = trans (drop-mid′ p₁ ws h refl refl) (drop-mid′ p₂ h xs refl refl) +-} + ------------------------------------------------------------------------ -- DEPRECATED NAMES ------------------------------------------------------------------------ @@ -218,5 +220,3 @@ shift v xs ys = ↭-shift {v = v} xs {ys} Please use ↭-shift instead. \\ NB. Parametrisation has changed to make v and ys implicit." #-} - - diff --git a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda index 3e91679bb2..a4c51ef423 100644 --- a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda @@ -4,7 +4,7 @@ -- Properties of the sublist relation over setoid equality. ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible #-} open import Relation.Binary.Definitions hiding (Decidable) open import Relation.Binary.Structures using (IsPreorder) From 439ef1916b90bba93b0c556e1bab7e06f67fc7c9 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 12 Mar 2024 18:45:20 +0000 Subject: [PATCH 31/65] involutive properties of symmetry, at all levels --- .../Binary/Permutation/Homogeneous.agda | 27 ++++++++++++++++++- .../Binary/Permutation/Propositional.agda | 12 ++++----- .../Permutation/Propositional/Properties.agda | 27 +++++++------------ .../Binary/Permutation/Setoid/Properties.agda | 20 +++++++++++--- .../PropositionalEquality/Properties.agda | 3 +++ 5 files changed, 61 insertions(+), 28 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index e9aff5ba28..9c6e6335f1 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -98,10 +98,13 @@ module _ {R : Rel A r} where sym R-sym (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (sym R-sym xs↭ys) sym R-sym (trans xs↭ys ys↭zs) = trans (sym R-sym ys↭zs) (sym R-sym xs↭ys) + ↭-sym′ : Symmetric R → Symmetric (Permutation R) + ↭-sym′ = sym + isEquivalence : Reflexive R → Symmetric R → IsEquivalence (Permutation R) isEquivalence R-refl R-sym = record { refl = ↭-refl′ R-refl - ; sym = sym R-sym + ; sym = ↭-sym′ R-sym ; trans = trans } @@ -124,6 +127,28 @@ module _ {R : Rel A r} ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs +------------------------------------------------------------------------ +-- A higher-dimensional property useful for the `Propositional` case + +module _ {R : Rel A r} {R-sym : Symmetric R} + (R-sym-involutive : ∀ {x y} → (p : R x y) → R-sym (R-sym p) ≡ p) + where + + ≋-sym-involutive′ : (p : Pointwise R xs ys) → + Pointwise.symmetric R-sym (Pointwise.symmetric R-sym p) ≡ p + ≋-sym-involutive′ [] = ≡.refl + ≋-sym-involutive′ (x∼y ∷ xs≋ys) + rewrite R-sym-involutive x∼y = ≡.cong (_ ∷_) (≋-sym-involutive′ xs≋ys) + + ↭-sym-involutive′ : (p : Permutation R xs ys) → ↭-sym′ R-sym (↭-sym′ R-sym p) ≡ p + ↭-sym-involutive′ (refl xs≋ys) = ≡.cong refl (≋-sym-involutive′ xs≋ys) + ↭-sym-involutive′ (prep eq p) = ≡.cong₂ prep (R-sym-involutive eq) (↭-sym-involutive′ p) + ↭-sym-involutive′ (swap eq₁ eq₂ p) + rewrite R-sym-involutive eq₁ | R-sym-involutive eq₂ + = ≡.cong (swap _ _) (↭-sym-involutive′ p) + ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) + + ------------------------------------------------------------------------ -- Map diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index da092eb019..5efbbe76a1 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -43,14 +43,14 @@ open ↭ public ↭-pointwise xs≋ys = ↭-reflexive (≋⇒≡ xs≋ys) -- A smart version of trans that avoids unnecessary `refl`s (see #1113). +↭-transˡ-≋ : LeftTrans _≋_ _↭_ +↭-transˡ-≋ xs≋ys ys↭zs with refl ← ≋⇒≡ xs≋ys = ys↭zs + +↭-transʳ-≋ : RightTrans _↭_ _≋_ +↭-transʳ-≋ xs↭ys ys≋zs with refl ← ≋⇒≡ ys≋zs = xs↭ys + ↭-trans : Transitive _↭_ ↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ - where - ↭-transˡ-≋ : LeftTrans _≋_ _↭_ - ↭-transˡ-≋ xs≋ys ys↭zs with refl ← ≋⇒≡ xs≋ys = ys↭zs - - ↭-transʳ-≋ : RightTrans _↭_ _≋_ - ↭-transʳ-≋ xs↭ys ys≋zs with refl ← ≋⇒≡ ys≋zs = xs↭ys ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 4854064665..be8940b224 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -12,25 +12,22 @@ open import Algebra.Bundles open import Algebra.Definitions open import Algebra.Structures open import Data.Bool.Base using (Bool; true; false) -open import Data.Nat using (suc) open import Data.Product.Base using (-,_; proj₂) open import Data.List.Base as List using (List; []; _∷_; [_]; _++_) -open import Data.List.Relation.Unary.Any using (Any; here; there) -open import Data.List.Relation.Unary.All using (All; []; _∷_) -open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) -import Data.List.Relation.Unary.Unique.Setoid as Unique open import Data.List.Membership.Propositional open import Data.List.Membership.Propositional.Properties import Data.List.Properties as List +open import Data.List.Relation.Unary.Any using (Any; here; there) +open import Data.List.Relation.Unary.All using (All; []; _∷_) open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) open import Function.Base using (_∘_; _⟨_⟩_) open import Level using (Level) -open import Relation.Unary using (Pred) open import Relation.Binary.Core using (Rel; _Preserves₂_⟶_⟶_) open import Relation.Binary.Definitions using (_Respects_; Decidable) open import Relation.Binary.PropositionalEquality as ≡ - using (_≡_ ; refl ; cong; cong₂; _≢_; setoid) + using (_≡_ ; refl ; cong; setoid) open import Relation.Nullary +open import Relation.Unary using (Pred) import Data.List.Relation.Binary.Permutation.Propositional as Propositional import Data.List.Relation.Binary.Permutation.Setoid.Properties as Permutation @@ -58,8 +55,10 @@ open ↭ public renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) -- DEPRECATION: legacy variation in implicit/explicit parametrisation hiding (shift) --- needing to specialise to ≡, where `Respects` and `Preserves` are trivial - hiding (map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭) +-- more efficient versions defined in `Propositional` + hiding (↭-transˡ-≋; ↭-transʳ-≋) +-- needing to specialise to ≡, where `Respects` and `Preserves` etc. are trivial + hiding (map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭; ↭-sym-involutive) ------------------------------------------------------------------------ -- Additional/specialised properties which hold in the case _≈_ = _≡_ @@ -75,14 +74,8 @@ open ↭ public -- sym ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p -↭-sym-involutive p = {!p!} -{- -↭-sym-involutive refl = refl -↭-sym-involutive (prep x ↭) = cong (prep x) (↭-sym-involutive ↭) -↭-sym-involutive (swap x y ↭) = cong (swap x y) (↭-sym-involutive ↭) -↭-sym-involutive (trans ↭₁ ↭₂) = - cong₂ trans (↭-sym-involutive ↭₁) (↭-sym-involutive ↭₂) --} +↭-sym-involutive = ↭.↭-sym-involutive {!≡.sym-involutive!} + ------------------------------------------------------------------------ -- Relationships to other predicates ------------------------------------------------------------------------ diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 76c74d69d0..983e57a45e 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -91,8 +91,20 @@ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ ↭-respˡ-≋ : _↭_ Respectsˡ _≋_ ↭-respˡ-≋ = Properties.↭-respˡ-≋ ≈-sym ≈-trans +↭-transˡ-≋ : LeftTrans _≋_ _↭_ +↭-transˡ-≋ = Properties.↭-transˡ-≋ ≈-trans + +↭-transʳ-≋ : RightTrans _↭_ _≋_ +↭-transʳ-≋ = Properties.↭-transʳ-≋ ≈-trans + +module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) + where + + ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p + ↭-sym-involutive = Properties.↭-sym-involutive′ ≈-sym-involutive + ------------------------------------------------------------------------ --- Properties of steps +-- Properties of steps (legacy) ------------------------------------------------------------------------ 0 Date: Tue, 12 Mar 2024 19:37:08 +0000 Subject: [PATCH 32/65] restores `--safe` while I ponder --- src/Data/List/Relation/Binary/BagAndSetEquality.agda | 6 +++--- .../Relation/Binary/Permutation/Homogeneous.agda | 5 ++++- .../Binary/Permutation/Propositional/Properties.agda | 12 ++++++------ .../Binary/Subset/Propositional/Properties.agda | 2 +- 4 files changed, 14 insertions(+), 11 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 9d4a50499e..bbc4c7910b 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -4,7 +4,7 @@ -- Bag and set equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible #-} +{-# OPTIONS --cubical-compatible --safe #-} module Data.List.Relation.Binary.BagAndSetEquality where @@ -556,7 +556,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ------------------------------------------------------------------------ -- Relationships to other relations - +{- ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ ↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where @@ -581,7 +581,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = -} to∘from : (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q to∘from p with res ← from∘to (↭-sym p) rewrite ↭-sym-involutive p = res - +-} ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl ∼bag⇒↭ {A = A} {x ∷ xs} eq diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 9c6e6335f1..f3f9632047 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -180,7 +180,10 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where Any-resp-↭ (swap x y p) (there (here px)) = here (resp y px) Any-resp-↭ (swap x y p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) - +{- + Any-resp-↭-here : ∀ x≈y xs↭ys px → Any-resp-↭ (x≈y ∷ xs↭ys) (here px) ≡ here (resp x≈y px) + Any-resp-↭-here xs↭ys px = ? +-} module _ {R : Rel A r} {S : Rel A s} (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index be8940b224..ded7a00e70 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -4,7 +4,7 @@ -- Properties of permutation in the Propositional case ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --allow-unsolved-metas #-} +{-# OPTIONS --cubical-compatible --safe #-} module Data.List.Relation.Binary.Permutation.Propositional.Properties {a} {A : Set a} where @@ -74,7 +74,7 @@ open ↭ public -- sym ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p -↭-sym-involutive = ↭.↭-sym-involutive {!≡.sym-involutive!} +↭-sym-involutive = ↭.↭-sym-involutive ≡.sym-involutive ------------------------------------------------------------------------ -- Relationships to other predicates @@ -88,11 +88,11 @@ Any-resp-↭ = ↭.Any-resp-↭ (≡.resp _) ∈-resp-↭ : (x ∈_) Respects _↭_ ∈-resp-↭ = Any-resp-↭ - +{- Any-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : Any P xs) → Any-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix Any-resp-[σ⁻¹∘σ] p ix = {!p!} -{- + Any-resp-[σ⁻¹∘σ] refl ix = refl Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl @@ -107,11 +107,11 @@ Any-resp-[σ⁻¹∘σ] (prep _ σ) (there ix) Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) rewrite Any-resp-[σ⁻¹∘σ] σ ix = refl --} + ∈-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : x ∈ xs) → ∈-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix ∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] - +-} ------------------------------------------------------------------------ -- map diff --git a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda index a4c51ef423..3e91679bb2 100644 --- a/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Subset/Propositional/Properties.agda @@ -4,7 +4,7 @@ -- Properties of the sublist relation over setoid equality. ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible #-} +{-# OPTIONS --cubical-compatible --safe #-} open import Relation.Binary.Definitions hiding (Decidable) open import Relation.Binary.Structures using (IsPreorder) From 6a7088fde673c9888f3d5294ec0c35f01b5b2ea3 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 06:20:04 +0000 Subject: [PATCH 33/65] use the import of equality primitives --- .../Relation/Binary/Permutation/Propositional/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index ded7a00e70..dca6e191fc 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -120,7 +120,7 @@ module _ {B : Set b} (f : A → B) where open Propositional {A = B} using () renaming (_↭_ to _↭′_) map⁺ : xs ↭ ys → List.map f xs ↭′ List.map f ys - map⁺ = ↭.map⁺ (setoid B) (≡.cong f) + map⁺ = ↭.map⁺ (setoid B) (cong f) {- -- permutations preserve 'being a mapped list' ↭-map-inv : List.map f xs ↭′ vs → ∃ λ ys → vs ≡ List.map f ys × xs ↭ ys From 53e7eb4e764ecffee3924bd8b5c390a8e43a159f Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 12:30:31 +0000 Subject: [PATCH 34/65] refactored construction --- .../Relation/Binary/BagAndSetEquality.agda | 33 ++++++++----------- 1 file changed, 13 insertions(+), 20 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index bbc4c7910b..ba090967dc 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -4,7 +4,7 @@ -- Bag and set equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible #-} module Data.List.Relation.Binary.BagAndSetEquality where @@ -30,7 +30,7 @@ open import Data.Sum.Properties hiding (map-cong) open import Data.Sum.Function.Propositional using (_⊎-cong_) open import Data.Unit.Polymorphic.Base open import Function.Base -open import Function.Bundles using (_↔_; Inverse; Equivalence; mk↔ₛ′; mk⇔) +open import Function.Bundles using (_↔_; Inverse; Equivalence; mk↔ₛ′; mk↔; mk⇔) open import Function.Related.Propositional as Related using (↔⇒; ⌊_⌋; ⌊_⌋→; ⇒→; K-refl; SK-sym) open import Function.Related.TypeIsomorphisms @@ -556,9 +556,10 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = ------------------------------------------------------------------------ -- Relationships to other relations -{- + ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ -↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) +↭⇒∼bag {A = A} xs↭ys {v} = mk↔ {to = to xs↭ys} {from = from xs↭ys} (from∘to xs↭ys , to∘from xs↭ys) + --mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where to : xs ↭ ys → v ∈ xs → v ∈ ys to = ∈-resp-↭ @@ -566,22 +567,14 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = from : xs ↭ ys → v ∈ ys → v ∈ xs from = ∈-resp-↭ ∘ ↭-sym - from∘to : (p : xs ↭ ys) (q : v ∈ xs) → from p (to p q) ≡ q - from∘to p v∈xs = {!p!} -{- - from∘to refl v∈xs = refl - from∘to (prep _ _) (here refl) = refl - from∘to (prep _ p) (there v∈xs) = ≡.cong there (from∘to p v∈xs) - from∘to (swap x y p) (here refl) = refl - from∘to (swap x y p) (there (here refl)) = refl - from∘to (swap x y p) (there (there v∈xs)) = ≡.cong (there ∘ there) (from∘to p v∈xs) - from∘to (trans p₁ p₂) v∈xs - rewrite from∘to p₂ (∈-resp-↭ p₁ v∈xs) - | from∘to p₁ v∈xs = refl --} - to∘from : (p : xs ↭ ys) (q : v ∈ ys) → to p (from p q) ≡ q - to∘from p with res ← from∘to (↭-sym p) rewrite ↭-sym-involutive p = res --} + from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → + ix ≡ from p iy → to p ix ≡ iy + from∘to = ∈-resp-↭-from∘to + + to∘from : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → + ix ≡ to p iy → from p ix ≡ iy + to∘from = ∈-resp-↭-to∘from + ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl ∼bag⇒↭ {A = A} {x ∷ xs} eq From c7e606c0d9d7adfcec4f35eec112e6b27a5ac49b Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 13:50:48 +0000 Subject: [PATCH 35/65] interim commit: one goal remaining! --- .../Relation/Binary/BagAndSetEquality.agda | 4 +- .../Binary/Permutation/Homogeneous.agda | 102 +++++++++++++----- .../Binary/Permutation/Propositional.agda | 2 +- .../Permutation/Propositional/Properties.agda | 37 ++----- .../Relation/Binary/Permutation/Setoid.agda | 2 +- .../Binary/Permutation/Setoid/Properties.agda | 19 +++- 6 files changed, 108 insertions(+), 58 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index ba090967dc..731d394473 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -4,7 +4,7 @@ -- Bag and set equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible #-} +{-# OPTIONS --cubical-compatible --safe #-} module Data.List.Relation.Binary.BagAndSetEquality where @@ -558,7 +558,7 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = -- Relationships to other relations ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ -↭⇒∼bag {A = A} xs↭ys {v} = mk↔ {to = to xs↭ys} {from = from xs↭ys} (from∘to xs↭ys , to∘from xs↭ys) +↭⇒∼bag {A = A} xs↭ys {v} = mk↔ (from∘to xs↭ys , to∘from xs↭ys) --mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where to : xs ↭ ys → v ∈ xs → v ∈ ys diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index f3f9632047..258ab44011 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -4,7 +4,7 @@ -- A definition for the permutation relation using setoid equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible --allow-unsolved-metas #-} module Data.List.Relation.Binary.Permutation.Homogeneous where @@ -127,28 +127,6 @@ module _ {R : Rel A r} ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs ------------------------------------------------------------------------- --- A higher-dimensional property useful for the `Propositional` case - -module _ {R : Rel A r} {R-sym : Symmetric R} - (R-sym-involutive : ∀ {x y} → (p : R x y) → R-sym (R-sym p) ≡ p) - where - - ≋-sym-involutive′ : (p : Pointwise R xs ys) → - Pointwise.symmetric R-sym (Pointwise.symmetric R-sym p) ≡ p - ≋-sym-involutive′ [] = ≡.refl - ≋-sym-involutive′ (x∼y ∷ xs≋ys) - rewrite R-sym-involutive x∼y = ≡.cong (_ ∷_) (≋-sym-involutive′ xs≋ys) - - ↭-sym-involutive′ : (p : Permutation R xs ys) → ↭-sym′ R-sym (↭-sym′ R-sym p) ≡ p - ↭-sym-involutive′ (refl xs≋ys) = ≡.cong refl (≋-sym-involutive′ xs≋ys) - ↭-sym-involutive′ (prep eq p) = ≡.cong₂ prep (R-sym-involutive eq) (↭-sym-involutive′ p) - ↭-sym-involutive′ (swap eq₁ eq₂ p) - rewrite R-sym-involutive eq₁ | R-sym-involutive eq₂ - = ≡.cong (swap _ _) (↭-sym-involutive′ p) - ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) - - ------------------------------------------------------------------------ -- Map @@ -180,10 +158,84 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where Any-resp-↭ (swap x y p) (there (here px)) = here (resp y px) Any-resp-↭ (swap x y p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) + {- - Any-resp-↭-here : ∀ x≈y xs↭ys px → Any-resp-↭ (x≈y ∷ xs↭ys) (here px) ≡ here (resp x≈y px) - Any-resp-↭-here xs↭ys px = ? + module _ {R-refl : Reflexive R} {R-sym : Symmetric R} {R-trans : Transitive R} + (R-sym-involutive : ∀ {x y} (p : R x y) → R-sym (R-sym p) ≡ p) + (R-trans-sym : ∀ {x y} (p : R x y) → R-trans p (R-sym p) ≡ R-refl {x = x}) + (R-resp-refl : ∀ {z} (pz : P z) → resp R-refl pz ≡ pz) + (R-resp-trans : ∀ {x y z} (p : R x y) (q : R y z) (px : P x) → + resp (R-trans p q) px ≡ resp q (resp p px)) + where + + Any-resp-↭-trans-sym : ∀ {xs ys} (xs↭ys : Permutation R xs ys) px → + Any-resp-↭ (↭-sym′ R-sym xs↭ys) (Any-resp-↭ xs↭ys px) ≡ px + Any-resp-↭-trans-sym (refl xs≋ys) pxs = {!!} + Any-resp-↭-trans-sym (prep x≈y xs↭ys) (here px) + rewrite ≡.sym (R-resp-trans x≈y (R-sym x≈y) px) + | R-trans-sym x≈y + | R-resp-refl px + = ≡.refl + Any-resp-↭-trans-sym (prep x≈y xs↭ys) (there pys) = {!!} + Any-resp-↭-trans-sym (swap eq₁ eq₂ xs↭ys) px = {!!} + Any-resp-↭-trans-sym (trans xs↭ys ys↭zs) px = {!!} -} + +------------------------------------------------------------------------ +-- A higher-dimensional property useful for the `Propositional` case +-- used in the equivalence proof between `Bag` equality and `_↭_` + + +------------------------------------------------------------------------ +-- Two higher-dimensional properties useful in the `Propositional` case, +-- specifically in the equivalence proof between `Bag` equality and `_↭_` + +module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where + + open IsEquivalence isEquivalence + renaming (refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) + + ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) + ∈-resp-↭ (refl xs≋ys) ixs = ∈-resp-Pointwise xs≋ys ixs + where + ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) + ∈-resp-Pointwise (x≈y ∷ xs) (here ix) = here (≈-trans ix x≈y) + ∈-resp-Pointwise (x≈y ∷ xs) (there ixs) = there (∈-resp-Pointwise xs ixs) + ∈-resp-↭ (prep x≈y p) (here ix) = here (≈-trans ix x≈y) + ∈-resp-↭ (prep x≈y p) (there ixs) = there (∈-resp-↭ p ixs) + ∈-resp-↭ (swap x y p) (here ix) = there (here (≈-trans ix x)) + ∈-resp-↭ (swap x y p) (there (here ix)) = here (≈-trans ix y) + ∈-resp-↭ (swap x y p) (there (there ixs)) = there (there (∈-resp-↭ p ixs)) + ∈-resp-↭ (trans p₁ p₂) ixs = ∈-resp-↭ p₂ (∈-resp-↭ p₁ ixs) + + module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) where + + ↭-sym-involutive′ : (p : Permutation _≈_ xs ys) → ↭-sym′ ≈-sym (↭-sym′ ≈-sym p) ≡ p + ↭-sym-involutive′ (refl xs≋ys) = ≡.cong refl (≋-sym-involutive′ xs≋ys) + where + ≋-sym-involutive′ : (p : Pointwise _≈_ xs ys) → + Pointwise.symmetric ≈-sym (Pointwise.symmetric ≈-sym p) ≡ p + ≋-sym-involutive′ [] = ≡.refl + ≋-sym-involutive′ (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y + = ≡.cong (_ ∷_) (≋-sym-involutive′ xs≋ys) + ↭-sym-involutive′ (prep eq p) = ≡.cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive′ p) + ↭-sym-involutive′ (swap eq₁ eq₂ p) rewrite ≈-sym-involutive eq₁ | ≈-sym-involutive eq₂ + = ≡.cong (swap _ _) (↭-sym-involutive′ p) + ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) + + ∈-resp-↭-sym : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym′ ≈-sym p) ix ≡ iy + ∈-resp-↭-sym p eq = {!!} + + ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → + ix ≡ ∈-resp-↭ (↭-sym′ ≈-sym p) iy → ∈-resp-↭ p ix ≡ iy + ∈-resp-↭-sym⁻¹ p eq + with eq′ ← ∈-resp-↭-sym (↭-sym′ ≈-sym p) rewrite ↭-sym-involutive′ p = eq′ eq + + +------------------------------------------------------------------------ +-- + module _ {R : Rel A r} {S : Rel A s} (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 5efbbe76a1..a37115a2d0 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -4,7 +4,7 @@ -- An inductive definition for the permutation relation ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible #-} module Data.List.Relation.Binary.Permutation.Propositional {a} {A : Set a} where diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index dca6e191fc..6a007d0b3f 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -36,7 +36,7 @@ private variable b p r : Level B : Set b - x y z : A + v x y z : A ws xs ys zs : List A vs : List B P : Pred A p @@ -87,31 +87,16 @@ Any-resp-↭ : (Any P) Respects _↭_ Any-resp-↭ = ↭.Any-resp-↭ (≡.resp _) ∈-resp-↭ : (x ∈_) Respects _↭_ -∈-resp-↭ = Any-resp-↭ -{- -Any-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : Any P xs) → - Any-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix -Any-resp-[σ⁻¹∘σ] p ix = {!p!} - -Any-resp-[σ⁻¹∘σ] refl ix = refl -Any-resp-[σ⁻¹∘σ] (prep _ _) (here _) = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ _) (here _) = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ _) (there (here _)) = refl -Any-resp-[σ⁻¹∘σ] (trans σ₁ σ₂) ix - rewrite Any-resp-[σ⁻¹∘σ] σ₂ (Any-resp-↭ σ₁ ix) - rewrite Any-resp-[σ⁻¹∘σ] σ₁ ix - = refl -Any-resp-[σ⁻¹∘σ] (prep _ σ) (there ix) - rewrite Any-resp-[σ⁻¹∘σ] σ ix - = refl -Any-resp-[σ⁻¹∘σ] (swap _ _ σ) (there (there ix)) - rewrite Any-resp-[σ⁻¹∘σ] σ ix - = refl - -∈-resp-[σ⁻¹∘σ] : (σ : xs ↭ ys) (ix : x ∈ xs) → - ∈-resp-↭ (↭-trans σ (↭-sym σ)) ix ≡ ix -∈-resp-[σ⁻¹∘σ] = Any-resp-[σ⁻¹∘σ] --} +∈-resp-↭ = ↭.∈-resp-↭ + +∈-resp-↭-from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → + ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy +∈-resp-↭-from∘to p = ↭.∈-resp-↭-sym⁻¹ ≡.sym-involutive p + +∈-resp-↭-to∘from : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy +∈-resp-↭-to∘from p = ↭.∈-resp-↭-sym ≡.sym-involutive p + ------------------------------------------------------------------------ -- map diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 14210b09ab..8c8de4eb9b 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -4,7 +4,7 @@ -- A definition for the permutation relation using setoid equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --cubical-compatible -safe #-} open import Function.Base using (_∘′_) open import Relation.Binary.Core using (Rel; _⇒_) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 983e57a45e..4584676205 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -43,7 +43,7 @@ import Data.List.Relation.Binary.Permutation.Setoid as Permutation import Data.List.Relation.Binary.Permutation.Homogeneous as Properties open Setoid S - using (_≈_) + using (_≈_; isEquivalence) renaming (Carrier to A; refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) open Permutation S open Membership S @@ -73,7 +73,7 @@ AllPairs-resp-↭ : Symmetric R → R Respects₂ _≈_ → (AllPairs R) Respect AllPairs-resp-↭ = Properties.AllPairs-resp-↭ ∈-resp-↭ : (x ∈_) Respects _↭_ -∈-resp-↭ = Any-resp-↭ (flip ≈-trans) +∈-resp-↭ = Properties.∈-resp-↭ isEquivalence Unique-resp-↭ : Unique Respects _↭_ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ @@ -101,7 +101,20 @@ module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym where ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p - ↭-sym-involutive = Properties.↭-sym-involutive′ ≈-sym-involutive + ↭-sym-involutive = Properties.↭-sym-involutive′ isEquivalence ≈-sym-involutive + + module _ {-(≈-trans-sym : ∀ {x y} (p : x ≈ y) → ≈-trans p (≈-sym p) ≡ ≈-refl {x = x}) + (≈-resp-refl : ∀ {x z} (p : x ≈ z) → ≈-trans p ≈-refl ≡ p) + (≈-resp-trans : ∀ {w x y z} (p : x ≈ y) (q : y ≈ z) (r : w ≈ x) → + ≈-trans r (≈-trans p q) ≡ ≈-trans (≈-trans r p) q)-} + where + + ∈-resp-↭-sym⁻¹ : ∀ (p : xs ↭ ys) {ix : x ∈ xs} {iy : x ∈ ys} → + ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy + ∈-resp-↭-sym⁻¹ p = Properties.∈-resp-↭-sym⁻¹ isEquivalence ≈-sym-involutive p + ∈-resp-↭-sym : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy + ∈-resp-↭-sym p = Properties.∈-resp-↭-sym isEquivalence ≈-sym-involutive p ------------------------------------------------------------------------ -- Properties of steps (legacy) From ec9d0f830997ca8df5de3a8a75b99c17d39e1313 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 15:13:45 +0000 Subject: [PATCH 36/65] `Homogeneous` safe! --- .../Binary/Permutation/Homogeneous.agda | 53 +++++++++++++++---- 1 file changed, 42 insertions(+), 11 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 258ab44011..eb72fb1113 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -4,7 +4,7 @@ -- A definition for the permutation relation using setoid equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible --allow-unsolved-metas #-} +{-# OPTIONS --cubical-compatible --safe #-} module Data.List.Relation.Binary.Permutation.Homogeneous where @@ -195,12 +195,12 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where open IsEquivalence isEquivalence renaming (refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) + ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) + ∈-resp-Pointwise (x≈y ∷ xs) (here ix) = here (≈-trans ix x≈y) + ∈-resp-Pointwise (x≈y ∷ xs) (there ixs) = there (∈-resp-Pointwise xs ixs) + ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) ∈-resp-↭ (refl xs≋ys) ixs = ∈-resp-Pointwise xs≋ys ixs - where - ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) - ∈-resp-Pointwise (x≈y ∷ xs) (here ix) = here (≈-trans ix x≈y) - ∈-resp-Pointwise (x≈y ∷ xs) (there ixs) = there (∈-resp-Pointwise xs ixs) ∈-resp-↭ (prep x≈y p) (here ix) = here (≈-trans ix x≈y) ∈-resp-↭ (prep x≈y p) (there ixs) = there (∈-resp-↭ p ixs) ∈-resp-↭ (swap x y p) (here ix) = there (here (≈-trans ix x)) @@ -208,13 +208,22 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where ∈-resp-↭ (swap x y p) (there (there ixs)) = there (there (∈-resp-↭ p ixs)) ∈-resp-↭ (trans p₁ p₂) ixs = ∈-resp-↭ p₂ (∈-resp-↭ p₁ ixs) - module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) where + module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) + (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → + ≈-trans (≈-trans p q) (≈-sym q) ≡ p) + + where + + private + ≋-sym : Symmetric (Pointwise _≈_) + ≋-sym = Pointwise.symmetric ≈-sym + ↭-sym : Symmetric (Permutation _≈_) + ↭-sym = ↭-sym′ ≈-sym - ↭-sym-involutive′ : (p : Permutation _≈_ xs ys) → ↭-sym′ ≈-sym (↭-sym′ ≈-sym p) ≡ p + ↭-sym-involutive′ : (p : Permutation _≈_ xs ys) → ↭-sym (↭-sym p) ≡ p ↭-sym-involutive′ (refl xs≋ys) = ≡.cong refl (≋-sym-involutive′ xs≋ys) where - ≋-sym-involutive′ : (p : Pointwise _≈_ xs ys) → - Pointwise.symmetric ≈-sym (Pointwise.symmetric ≈-sym p) ≡ p + ≋-sym-involutive′ : (p : Pointwise _≈_ xs ys) → ≋-sym (≋-sym p) ≡ p ≋-sym-involutive′ [] = ≡.refl ≋-sym-involutive′ (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y = ≡.cong (_ ∷_) (≋-sym-involutive′ xs≋ys) @@ -223,9 +232,31 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where = ≡.cong (swap _ _) (↭-sym-involutive′ p) ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) + ∈-resp-Pointwise-sym : (p : Pointwise _≈_ ys xs) → + {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ix ≡ ∈-resp-Pointwise p iy → + ∈-resp-Pointwise (≋-sym p) ix ≡ iy + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} {here iy} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym ix x≈y) + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} {there iys} eq + with ≡.refl ← eq = cong there (∈-resp-Pointwise-sym xs≋ys ≡.refl) + ∈-resp-↭-sym : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym′ ≈-sym p) ix ≡ iy - ∈-resp-↭-sym p eq = {!!} + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy + ∈-resp-↭-sym (refl xs≋ys) eq = ∈-resp-Pointwise-sym xs≋ys eq + ∈-resp-↭-sym (prep eq₁ p) {here iy} {here ix} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) + ∈-resp-↭-sym (prep eq₁ p) {there iys} {there ixs} eq + with ≡.refl ← eq = cong there (∈-resp-↭-sym p ≡.refl) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {here iy} () + ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {there iys} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq + with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq + with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym p ≡.refl) + ∈-resp-↭-sym (trans p₁ p₂) eq = ∈-resp-↭-sym p₁ (∈-resp-↭-sym p₂ eq) ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → ix ≡ ∈-resp-↭ (↭-sym′ ≈-sym p) iy → ∈-resp-↭ p ix ≡ iy From 0c0828e66e30bc9a45e9e95f642e2c3b4908e9fc Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 15:23:31 +0000 Subject: [PATCH 37/65] `Setoid` safe! --- .../Binary/Permutation/Homogeneous.agda | 94 +++++++------------ .../Relation/Binary/Permutation/Setoid.agda | 2 +- .../Binary/Permutation/Setoid/Properties.agda | 17 ++-- 3 files changed, 45 insertions(+), 68 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index eb72fb1113..0aeab7b9c1 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -159,33 +159,6 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where Any-resp-↭ (swap x y p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) -{- - module _ {R-refl : Reflexive R} {R-sym : Symmetric R} {R-trans : Transitive R} - (R-sym-involutive : ∀ {x y} (p : R x y) → R-sym (R-sym p) ≡ p) - (R-trans-sym : ∀ {x y} (p : R x y) → R-trans p (R-sym p) ≡ R-refl {x = x}) - (R-resp-refl : ∀ {z} (pz : P z) → resp R-refl pz ≡ pz) - (R-resp-trans : ∀ {x y z} (p : R x y) (q : R y z) (px : P x) → - resp (R-trans p q) px ≡ resp q (resp p px)) - where - - Any-resp-↭-trans-sym : ∀ {xs ys} (xs↭ys : Permutation R xs ys) px → - Any-resp-↭ (↭-sym′ R-sym xs↭ys) (Any-resp-↭ xs↭ys px) ≡ px - Any-resp-↭-trans-sym (refl xs≋ys) pxs = {!!} - Any-resp-↭-trans-sym (prep x≈y xs↭ys) (here px) - rewrite ≡.sym (R-resp-trans x≈y (R-sym x≈y) px) - | R-trans-sym x≈y - | R-resp-refl px - = ≡.refl - Any-resp-↭-trans-sym (prep x≈y xs↭ys) (there pys) = {!!} - Any-resp-↭-trans-sym (swap eq₁ eq₂ xs↭ys) px = {!!} - Any-resp-↭-trans-sym (trans xs↭ys ys↭zs) px = {!!} --} - ------------------------------------------------------------------------- --- A higher-dimensional property useful for the `Propositional` case --- used in the equivalence proof between `Bag` equality and `_↭_` - - ------------------------------------------------------------------------ -- Two higher-dimensional properties useful in the `Propositional` case, -- specifically in the equivalence proof between `Bag` equality and `_↭_` @@ -209,8 +182,6 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where ∈-resp-↭ (trans p₁ p₂) ixs = ∈-resp-↭ p₂ (∈-resp-↭ p₁ ixs) module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) - (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → - ≈-trans (≈-trans p q) (≈-sym q) ≡ p) where @@ -232,36 +203,41 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where = ≡.cong (swap _ _) (↭-sym-involutive′ p) ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) - ∈-resp-Pointwise-sym : (p : Pointwise _≈_ ys xs) → - {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-Pointwise p iy → - ∈-resp-Pointwise (≋-sym p) ix ≡ iy - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} {here iy} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym ix x≈y) - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} {there iys} eq - with ≡.refl ← eq = cong there (∈-resp-Pointwise-sym xs≋ys ≡.refl) - - ∈-resp-↭-sym : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy - ∈-resp-↭-sym (refl xs≋ys) eq = ∈-resp-Pointwise-sym xs≋ys eq - ∈-resp-↭-sym (prep eq₁ p) {here iy} {here ix} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) - ∈-resp-↭-sym (prep eq₁ p) {there iys} {there ixs} eq - with ≡.refl ← eq = cong there (∈-resp-↭-sym p ≡.refl) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {here iy} () - ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {there iys} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq - with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq - with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym p ≡.refl) - ∈-resp-↭-sym (trans p₁ p₂) eq = ∈-resp-↭-sym p₁ (∈-resp-↭-sym p₂ eq) - - ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → - ix ≡ ∈-resp-↭ (↭-sym′ ≈-sym p) iy → ∈-resp-↭ p ix ≡ iy - ∈-resp-↭-sym⁻¹ p eq - with eq′ ← ∈-resp-↭-sym (↭-sym′ ≈-sym p) rewrite ↭-sym-involutive′ p = eq′ eq + module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → + ≈-trans (≈-trans p q) (≈-sym q) ≡ p) + + where + + ∈-resp-Pointwise-sym : (p : Pointwise _≈_ ys xs) → + {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ix ≡ ∈-resp-Pointwise p iy → + ∈-resp-Pointwise (≋-sym p) ix ≡ iy + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} {here iy} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym ix x≈y) + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} {there iys} eq + with ≡.refl ← eq = cong there (∈-resp-Pointwise-sym xs≋ys ≡.refl) + + ∈-resp-↭-sym : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy + ∈-resp-↭-sym (refl xs≋ys) eq = ∈-resp-Pointwise-sym xs≋ys eq + ∈-resp-↭-sym (prep eq₁ p) {here iy} {here ix} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) + ∈-resp-↭-sym (prep eq₁ p) {there iys} {there ixs} eq + with ≡.refl ← eq = cong there (∈-resp-↭-sym p ≡.refl) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {here iy} () + ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {there iys} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq + with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) + ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq + with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym p ≡.refl) + ∈-resp-↭-sym (trans p₁ p₂) eq = ∈-resp-↭-sym p₁ (∈-resp-↭-sym p₂ eq) + + ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → + ix ≡ ∈-resp-↭ (↭-sym′ ≈-sym p) iy → ∈-resp-↭ p ix ≡ iy + ∈-resp-↭-sym⁻¹ p eq + with eq′ ← ∈-resp-↭-sym (↭-sym′ ≈-sym p) rewrite ↭-sym-involutive′ p = eq′ eq ------------------------------------------------------------------------ diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 8c8de4eb9b..14210b09ab 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -4,7 +4,7 @@ -- A definition for the permutation relation using setoid equality ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible -safe #-} +{-# OPTIONS --cubical-compatible --safe #-} open import Function.Base using (_∘′_) open import Relation.Binary.Core using (Rel; _⇒_) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 4584676205..4475c33b1f 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -98,23 +98,24 @@ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ ↭-transʳ-≋ = Properties.↭-transʳ-≋ ≈-trans module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) - where + + where ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p ↭-sym-involutive = Properties.↭-sym-involutive′ isEquivalence ≈-sym-involutive - module _ {-(≈-trans-sym : ∀ {x y} (p : x ≈ y) → ≈-trans p (≈-sym p) ≡ ≈-refl {x = x}) - (≈-resp-refl : ∀ {x z} (p : x ≈ z) → ≈-trans p ≈-refl ≡ p) - (≈-resp-trans : ∀ {w x y z} (p : x ≈ y) (q : y ≈ z) (r : w ≈ x) → - ≈-trans r (≈-trans p q) ≡ ≈-trans (≈-trans r p) q)-} - where + module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → + ≈-trans (≈-trans p q) (≈-sym q) ≡ p) + where ∈-resp-↭-sym⁻¹ : ∀ (p : xs ↭ ys) {ix : x ∈ xs} {iy : x ∈ ys} → ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy - ∈-resp-↭-sym⁻¹ p = Properties.∈-resp-↭-sym⁻¹ isEquivalence ≈-sym-involutive p + ∈-resp-↭-sym⁻¹ p = Properties.∈-resp-↭-sym⁻¹ + isEquivalence ≈-sym-involutive ≈-trans-trans-sym p ∈-resp-↭-sym : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy - ∈-resp-↭-sym p = Properties.∈-resp-↭-sym isEquivalence ≈-sym-involutive p + ∈-resp-↭-sym p = Properties.∈-resp-↭-sym + isEquivalence ≈-sym-involutive ≈-trans-trans-sym p ------------------------------------------------------------------------ -- Properties of steps (legacy) From ca7607295885c8990aad9d25ab4e835fadbac871 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 15:41:00 +0000 Subject: [PATCH 38/65] `Propositional` safe! `equality` proofs reverted in favour of local defns --- .../Binary/Permutation/Propositional.agda | 2 +- .../Permutation/Propositional/Properties.agda | 21 ++++++++++++------- .../PropositionalEquality/Properties.agda | 3 --- 3 files changed, 15 insertions(+), 11 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index a37115a2d0..5efbbe76a1 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -4,7 +4,7 @@ -- An inductive definition for the permutation relation ------------------------------------------------------------------------ -{-# OPTIONS --cubical-compatible #-} +{-# OPTIONS --cubical-compatible --safe #-} module Data.List.Relation.Binary.Permutation.Propositional {a} {A : Set a} where diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 6a007d0b3f..243e358d99 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -51,19 +51,26 @@ private open Propositional {A = A} public open ↭ public --- legacy variations in naming +-- POSSIBLE DEPRECATION: legacy variations in naming renaming (dropMiddleElement-≋ to drop-mid-≡; dropMiddleElement to drop-mid) -- DEPRECATION: legacy variation in implicit/explicit parametrisation hiding (shift) -- more efficient versions defined in `Propositional` hiding (↭-transˡ-≋; ↭-transʳ-≋) -- needing to specialise to ≡, where `Respects` and `Preserves` etc. are trivial - hiding (map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭; ↭-sym-involutive) + hiding ( map⁺; All-resp-↭; Any-resp-↭; ∈-resp-↭; ↭-sym-involutive + ; ∈-resp-↭-sym⁻¹; ∈-resp-↭-sym) ------------------------------------------------------------------------ -- Additional/specialised properties which hold in the case _≈_ = _≡_ ------------------------------------------------------------------------ +sym-involutive : (p : x ≡ y) → ≡.sym (≡.sym p) ≡ p +sym-involutive refl = refl + +trans-trans-sym : (p : x ≡ y) (q : y ≡ z) → ≡.trans (≡.trans p q) (≡.sym q) ≡ p +trans-trans-sym refl refl = refl + ------------------------------------------------------------------------ -- Permutations of singleton lists @@ -74,7 +81,7 @@ open ↭ public -- sym ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p -↭-sym-involutive = ↭.↭-sym-involutive ≡.sym-involutive +↭-sym-involutive = ↭.↭-sym-involutive sym-involutive ------------------------------------------------------------------------ -- Relationships to other predicates @@ -89,13 +96,13 @@ Any-resp-↭ = ↭.Any-resp-↭ (≡.resp _) ∈-resp-↭ : (x ∈_) Respects _↭_ ∈-resp-↭ = ↭.∈-resp-↭ -∈-resp-↭-from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → +∈-resp-↭-sym⁻¹ : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy -∈-resp-↭-from∘to p = ↭.∈-resp-↭-sym⁻¹ ≡.sym-involutive p +∈-resp-↭-sym⁻¹ p = ↭.∈-resp-↭-sym⁻¹ sym-involutive trans-trans-sym p -∈-resp-↭-to∘from : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → +∈-resp-↭-sym : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy -∈-resp-↭-to∘from p = ↭.∈-resp-↭-sym ≡.sym-involutive p +∈-resp-↭-sym p = ↭.∈-resp-↭-sym sym-involutive trans-trans-sym p ------------------------------------------------------------------------ -- map diff --git a/src/Relation/Binary/PropositionalEquality/Properties.agda b/src/Relation/Binary/PropositionalEquality/Properties.agda index 733c8971a3..a420e45b64 100644 --- a/src/Relation/Binary/PropositionalEquality/Properties.agda +++ b/src/Relation/Binary/PropositionalEquality/Properties.agda @@ -93,9 +93,6 @@ cong-∘ : ∀ {x y : A} {f : B → C} {g : A → B} (p : x ≡ y) → cong (f ∘ g) p ≡ cong f (cong g p) cong-∘ refl = refl -sym-involutive : ∀ {x y : A} (p : x ≡ y) → sym (sym p) ≡ p -sym-involutive refl = refl - sym-cong : ∀ {x y : A} {f : A → B} (p : x ≡ y) → sym (cong f p) ≡ cong f (sym p) sym-cong refl = refl From 5ab893d45f6bff0613678b8f70d12576fc72cca3 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 15:43:41 +0000 Subject: [PATCH 39/65] tidy up `BagAndSetEquality` --- src/Data/List/Relation/Binary/BagAndSetEquality.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 731d394473..3c0a7421a6 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -569,11 +569,11 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → ix ≡ from p iy → to p ix ≡ iy - from∘to = ∈-resp-↭-from∘to + from∘to = ∈-resp-↭-sym⁻¹ to∘from : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → ix ≡ to p iy → from p ix ≡ iy - to∘from = ∈-resp-↭-to∘from + to∘from = ∈-resp-↭-sym ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl From afe32c75648d480d78d3e045fe75684f81f0d1f8 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 16:04:55 +0000 Subject: [PATCH 40/65] typo --- doc/README/Data/List/Relation/Binary/Permutation.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/README/Data/List/Relation/Binary/Permutation.agda b/doc/README/Data/List/Relation/Binary/Permutation.agda index eed81a3645..08084b8267 100644 --- a/doc/README/Data/List/Relation/Binary/Permutation.agda +++ b/doc/README/Data/List/Relation/Binary/Permutation.agda @@ -30,7 +30,7 @@ open import Data.List.Relation.Binary.Permutation.Propositional -- a permutation of itself, the second `↭-prep` says that if the -- heads of the lists are the same they can be skipped, the third -- `↭-swap` says that the first two elements of the lists can be --- swapped and the fourth `↭trans` says that permutation proofs +-- swapped and the fourth `↭-trans` says that permutation proofs -- can be chained transitively. -- For example a proof that two lists are a permutation of one From f98f043adce593fa4f7ac691b9c5b2eddb19e848 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 13 Mar 2024 17:07:44 +0000 Subject: [PATCH 41/65] bracketing fixes the bug; but surely that's not the point? --- doc/README/Data/List/Relation/Binary/Permutation.agda | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/doc/README/Data/List/Relation/Binary/Permutation.agda b/doc/README/Data/List/Relation/Binary/Permutation.agda index 08084b8267..5a640b4a7b 100644 --- a/doc/README/Data/List/Relation/Binary/Permutation.agda +++ b/doc/README/Data/List/Relation/Binary/Permutation.agda @@ -56,13 +56,13 @@ lem₂ = begin -- with specialised combinators `<⟨_⟩` and `<<⟨_⟩` to support the -- distinguished use of the `↭-prep ` and `↭-swap` steps, allowing -- the above proof to be recast in the following form: -{- + lem₂ᵣ : 1 ∷ 2 ∷ 3 ∷ [] ↭ 3 ∷ 1 ∷ 2 ∷ [] lem₂ᵣ = begin - 1 ∷ 2 ∷ 3 ∷ [] <⟨ swap 2 3 ↭-refl ⟩ - 1 ∷ 3 ∷ 2 ∷ [] <<⟨ ↭-refl ⟩ - 3 ∷ 1 ∷ 2 ∷ [] ∎ --} + 1 ∷ (2 ∷ (3 ∷ [])) <⟨ ↭-swap 2 3 ↭-refl ⟩ + 1 ∷ 3 ∷ (2 ∷ []) <<⟨ ↭-refl ⟩ + 3 ∷ 1 ∷ (2 ∷ []) ∎ + -- As might be expected, properties of the permutation relation may be -- found in: From eb904ff6a28ae05b91c54cd2ba650321b1280ef8 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Thu, 14 Mar 2024 17:17:33 +0000 Subject: [PATCH 42/65] cosmetics --- .../Relation/Binary/BagAndSetEquality.agda | 44 +++++++++++++------ 1 file changed, 30 insertions(+), 14 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 3c0a7421a6..2874f008b7 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -537,36 +537,52 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = lemma : ∀ {xs ys} (inv : x ∷ xs ∼[ bag ] x ∷ ys) {z} → Well-behaved (Inverse.to (∼→⊎↔⊎ inv {z})) - lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ with begin - zero ≡⟨⟩ - index-of {xs = x ∷ xs} (here p) ≡⟨⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ ≡.sym $ to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) ≡⟨ ≡.cong index-of $ - strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩ - index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩ - index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) ≡⟨ ≡.cong index-of $ ≡.sym $ strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩ - index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) ≡⟨⟩ - index-of {xs = x ∷ xs} (there z∈xs) ≡⟨⟩ - suc (index-of {xs = xs} z∈xs) ∎ + lemma {xs} inv {b = z∈xs} {a = p} {a′ = q} hyp₁ hyp₂ = case contra of λ () where open Inverse open ≡-Reasoning - ... | () + contra : zero ≡ suc (index-of {xs = xs} z∈xs) + contra = begin + zero + ≡⟨⟩ + index-of {xs = x ∷ xs} (here p) + ≡⟨⟩ + index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₁ p) + ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₁ p} hyp₂ ⟨ + index-of {xs = x ∷ xs} (to (∷↔ _) $ (from (∼→⊎↔⊎ inv) $ inj₁ q)) + ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here q)) ⟩ + index-of {xs = x ∷ xs} (to (SK-sym inv) $ here q) + ≡⟨ index-equality-preserved (SK-sym inv) refl ⟩ + index-of {xs = x ∷ xs} (to (SK-sym inv) $ here p) + ≡⟨ ≡.cong index-of $ strictlyInverseˡ (∷↔ _) (from inv (here p)) ⟨ + index-of {xs = x ∷ xs} (to (∷↔ _) (from (∼→⊎↔⊎ inv) $ inj₁ p)) + ≡⟨ ≡.cong (index-of ∘ (to (∷↔ (_ ≡_)))) $ to-from (∼→⊎↔⊎ inv) {x = inj₂ z∈xs} hyp₁ ⟩ + index-of {xs = x ∷ xs} (to (∷↔ _) $ inj₂ z∈xs) + ≡⟨⟩ + index-of {xs = x ∷ xs} (there z∈xs) + ≡⟨⟩ + suc (index-of {xs = xs} z∈xs) + ∎ ------------------------------------------------------------------------ -- Relationships to other relations ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ ↭⇒∼bag {A = A} xs↭ys {v} = mk↔ (from∘to xs↭ys , to∘from xs↭ys) - --mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) + -- mk↔ₛ′ (to xs↭ys) (from xs↭ys) {!to∘from xs↭ys!} {!from∘to xs↭ys!} where to : xs ↭ ys → v ∈ xs → v ∈ ys to = ∈-resp-↭ from : xs ↭ ys → v ∈ ys → v ∈ xs from = ∈-resp-↭ ∘ ↭-sym +{- + from∘to : (p : xs ↭ ys) {ix : v ∈ xs} → from p (to p ix) ≡ ix + from∘to p = ∈-resp-↭-sym p refl + to∘from : (p : ys ↭ xs) {ix : v ∈ xs} → to p (from p ix) ≡ ix + to∘from p = ∈-resp-↭-sym⁻¹ p refl +-} from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → ix ≡ from p iy → to p ix ≡ iy from∘to = ∈-resp-↭-sym⁻¹ From 3856163ddb3a5d03bbeea915cbc9d7b74a3c986c Mon Sep 17 00:00:00 2001 From: James McKinna Date: Thu, 14 Mar 2024 22:15:34 +0000 Subject: [PATCH 43/65] removed any need for `steps` or `Acc`-induction --- .../Binary/Permutation/Homogeneous.agda | 74 ++++++++++--------- 1 file changed, 41 insertions(+), 33 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 0aeab7b9c1..76096dba56 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -13,6 +13,7 @@ open import Data.List.Base as List using (List; []; _∷_; [_]) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +import Data.List.Relation.Unary.Any.Properties as Any open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) open import Data.Nat.Base using (ℕ; suc; _+_; _<_) @@ -24,7 +25,7 @@ open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_) open import Relation.Binary.Bundles using (Setoid) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning -open import Relation.Binary.Structures using (IsEquivalence) +open import Relation.Binary.Structures using (IsEquivalence; IsPartialEquivalence) open import Relation.Binary.Definitions using ( Reflexive; Symmetric; Transitive; LeftTrans; RightTrans ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) @@ -163,10 +164,10 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where -- Two higher-dimensional properties useful in the `Propositional` case, -- specifically in the equivalence proof between `Bag` equality and `_↭_` -module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where +module _ {_≈_ : Rel A r} (isPartialEquivalence : IsPartialEquivalence _≈_) where - open IsEquivalence isEquivalence - renaming (refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) + open IsPartialEquivalence isPartialEquivalence + renaming (sym to ≈-sym; trans to ≈-trans) ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) ∈-resp-Pointwise (x≈y ∷ xs) (here ix) = here (≈-trans ix x≈y) @@ -235,9 +236,9 @@ module _ {_≈_ : Rel A r} (isEquivalence : IsEquivalence _≈_) where ∈-resp-↭-sym (trans p₁ p₂) eq = ∈-resp-↭-sym p₁ (∈-resp-↭-sym p₂ eq) ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → - ix ≡ ∈-resp-↭ (↭-sym′ ≈-sym p) iy → ∈-resp-↭ p ix ≡ iy + ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy ∈-resp-↭-sym⁻¹ p eq - with eq′ ← ∈-resp-↭-sym (↭-sym′ ≈-sym p) rewrite ↭-sym-involutive′ p = eq′ eq + with eq′ ← ∈-resp-↭-sym (↭-sym p) rewrite ↭-sym-involutive′ p = eq′ eq ------------------------------------------------------------------------ @@ -274,8 +275,8 @@ module _ {R : Rel A r} where module _ {R : Rel A r} (R-refl : Reflexive R) where - ↭-prep : ∀ x {xs ys} → Permutation R xs ys → Permutation R (x ∷ xs) (x ∷ ys) - ↭-prep _ xs↭ys = prep R-refl xs↭ys + ↭-prep : ∀ {x} {xs ys} → Permutation R xs ys → Permutation R (x ∷ xs) (x ∷ ys) + ↭-prep xs↭ys = prep R-refl xs↭ys ↭-swap : ∀ x y {xs ys} → Permutation R xs ys → Permutation R (x ∷ y ∷ xs) (y ∷ x ∷ ys) ↭-swap _ _ xs↭ys = swap R-refl R-refl xs↭ys @@ -352,40 +353,47 @@ module _ {R : Rel A r} (R-sym : Symmetric R) (R-trans : Transitive R) where module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) where private - ≋-refl = Pointwise.refl {R = R} R-refl - ↭-refl = ↭-refl′ {R = R} R-refl + ≋-refl : Reflexive (Pointwise R) + ≋-refl = Pointwise.refl R-refl + ↭-refl : Reflexive (Permutation R) + ↭-refl = ↭-refl′ R-refl + ≋-trans : Transitive (Pointwise R) + ≋-trans = Pointwise.transitive R-trans _++[_]++_ = λ (xs : List A) z ys → xs List.++ List.[ z ] List.++ ys split-↭ : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) × Permutation R (ps List.++ qs) (as List.++ bs) - split-↭ v as bs p = helper as bs p (<-wellFounded (steps p)) + split-↭ v as bs p = -- no longer requires `Acc`-induction or `steps`... + helper as bs p ≋-refl where - helper : ∀ as bs {xs} (p : Permutation R xs (as ++[ v ]++ bs)) → - Acc _<_ (steps p) → + helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → + Pointwise R ys (as ++[ v ]++ bs) → ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) × Permutation R (ps List.++ qs) (as List.++ bs) - helper [] bs (refl eq) _ = [] , bs , eq , ↭-refl - helper (a ∷ []) bs (refl eq) _ = List.[ a ] , bs , eq , ↭-refl - helper (a ∷ b ∷ as) bs (refl eq) _ = a ∷ b ∷ as , bs , eq , ↭-refl - helper [] bs (prep v≈x xs↭bs) _ = [] , _ , v≈x ∷ ≋-refl , xs↭bs - helper (a ∷ as) bs (prep eq as↭xs) (acc rec) - with ps , qs , eq₂ , ↭ ← helper as bs as↭xs (rec (n<1+n _)) - = a ∷ ps , qs , eq ∷ eq₂ , prep R-refl ↭ - helper [] (b ∷ bs) (swap x≈b y≈v xs↭bs) _ - = List.[ b ] , _ , x≈b ∷ y≈v ∷ ≋-refl , prep R-refl xs↭bs - helper (a ∷ []) bs (swap x≈v y≈a xs↭bs) _ - = [] , a ∷ _ , x≈v ∷ y≈a ∷ ≋-refl , prep R-refl xs↭bs - helper (a ∷ b ∷ as) bs (swap x≈b y≈a as↭xs) (acc rec) - with ps , qs , eq , ↭ ← helper as bs as↭xs (rec (n<1+n _)) - = b ∷ a ∷ ps , qs , x≈b ∷ y≈a ∷ eq , swap R-refl R-refl ↭ - helper as bs (trans xs↭ys ys↭zs) (acc rec) - with ps , qs , eq , ↭ ← helper as bs ys↭zs (rec (m Date: Thu, 14 Mar 2024 22:19:19 +0000 Subject: [PATCH 44/65] knock-on changes from improvements to `Homogeneous` --- src/Data/List/Relation/Binary/Permutation/Setoid.agda | 2 +- .../Relation/Binary/Permutation/Setoid/Properties.agda | 10 +++++----- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 14210b09ab..7e831b9962 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -48,7 +48,7 @@ _↭_ = Homogeneous.Permutation _≈_ ↭-pointwise = Homogeneous.↭-pointwise ↭-prep : ∀ x {xs ys} → xs ↭ ys → x ∷ xs ↭ x ∷ ys -↭-prep = Homogeneous.↭-prep ≈-refl +↭-prep _ = Homogeneous.↭-prep ≈-refl ↭-swap : ∀ x y {xs ys} → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys ↭-swap = Homogeneous.↭-swap ≈-refl diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 4475c33b1f..1b03798252 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -43,7 +43,7 @@ import Data.List.Relation.Binary.Permutation.Setoid as Permutation import Data.List.Relation.Binary.Permutation.Homogeneous as Properties open Setoid S - using (_≈_; isEquivalence) + using (_≈_; isPartialEquivalence) renaming (Carrier to A; refl to ≈-refl; sym to ≈-sym; trans to ≈-trans) open Permutation S open Membership S @@ -73,7 +73,7 @@ AllPairs-resp-↭ : Symmetric R → R Respects₂ _≈_ → (AllPairs R) Respect AllPairs-resp-↭ = Properties.AllPairs-resp-↭ ∈-resp-↭ : (x ∈_) Respects _↭_ -∈-resp-↭ = Properties.∈-resp-↭ isEquivalence +∈-resp-↭ = Properties.∈-resp-↭ isPartialEquivalence Unique-resp-↭ : Unique Respects _↭_ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ @@ -102,7 +102,7 @@ module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym where ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p - ↭-sym-involutive = Properties.↭-sym-involutive′ isEquivalence ≈-sym-involutive + ↭-sym-involutive = Properties.↭-sym-involutive′ isPartialEquivalence ≈-sym-involutive module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → ≈-trans (≈-trans p q) (≈-sym q) ≡ p) @@ -111,11 +111,11 @@ module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym ∈-resp-↭-sym⁻¹ : ∀ (p : xs ↭ ys) {ix : x ∈ xs} {iy : x ∈ ys} → ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy ∈-resp-↭-sym⁻¹ p = Properties.∈-resp-↭-sym⁻¹ - isEquivalence ≈-sym-involutive ≈-trans-trans-sym p + isPartialEquivalence ≈-sym-involutive ≈-trans-trans-sym p ∈-resp-↭-sym : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy ∈-resp-↭-sym p = Properties.∈-resp-↭-sym - isEquivalence ≈-sym-involutive ≈-trans-trans-sym p + isPartialEquivalence ≈-sym-involutive ≈-trans-trans-sym p ------------------------------------------------------------------------ -- Properties of steps (legacy) From bf14b4d29eb421c4645f87e038f175fb854c14ac Mon Sep 17 00:00:00 2001 From: James McKinna Date: Thu, 14 Mar 2024 22:22:56 +0000 Subject: [PATCH 45/65] tidy up `BagAndSetEquality` --- .../Relation/Binary/BagAndSetEquality.agda | 24 +++++++------------ 1 file changed, 8 insertions(+), 16 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 2874f008b7..06f4df035a 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -568,28 +568,20 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = -- Relationships to other relations ↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_ -↭⇒∼bag {A = A} xs↭ys {v} = mk↔ (from∘to xs↭ys , to∘from xs↭ys) - -- mk↔ₛ′ (to xs↭ys) (from xs↭ys) {!to∘from xs↭ys!} {!from∘to xs↭ys!} +↭⇒∼bag {A = A} xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys) where + to : xs ↭ ys → v ∈ xs → v ∈ ys to = ∈-resp-↭ from : xs ↭ ys → v ∈ ys → v ∈ xs from = ∈-resp-↭ ∘ ↭-sym -{- - from∘to : (p : xs ↭ ys) {ix : v ∈ xs} → from p (to p ix) ≡ ix - from∘to p = ∈-resp-↭-sym p refl - - to∘from : (p : ys ↭ xs) {ix : v ∈ xs} → to p (from p ix) ≡ ix - to∘from p = ∈-resp-↭-sym⁻¹ p refl --} - from∘to : (p : xs ↭ ys) {ix : v ∈ xs} {iy : v ∈ ys} → - ix ≡ from p iy → to p ix ≡ iy - from∘to = ∈-resp-↭-sym⁻¹ - - to∘from : (p : ys ↭ xs) {iy : v ∈ ys} {ix : v ∈ xs} → - ix ≡ to p iy → from p ix ≡ iy - to∘from = ∈-resp-↭-sym + + from∘to : (p : xs ↭ ys) (ix : v ∈ xs) → from p (to p ix) ≡ ix + from∘to p _ = ∈-resp-↭-sym p refl + + to∘from : (p : xs ↭ ys) (iy : v ∈ ys) → to p (from p iy) ≡ iy + to∘from p _ = ∈-resp-↭-sym⁻¹ p refl ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl From 980909f4fd7516b6e3004d038349d4e0f3f6f3b5 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 15 Mar 2024 10:19:08 +0000 Subject: [PATCH 46/65] NB weird `Reasoning` problem/bug --- .../Relation/Binary/Permutation/Propositional/Properties.agda | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 243e358d99..c3b55dba9d 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -176,7 +176,9 @@ drop-mid {A = A} {x} ws xs p = drop-mid′ p ws xs refl refl _ ∷ _ ∷ xs ++ _ ∎ drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ []) refl refl = begin _ ∷ _ ∷ ws ++ _ <<⟨ refl ⟩ +-- *** NB the next step can't be replaced with <⟨ shift _ _ _ ⟨ for some reason? *** _ ∷ (_ ∷ ws ++ _) <⟨ ↭-sym (shift _ _ _) ⟩ +-- *** because of parsing problems for that use of the `Reasoning` combinators!? *** _ ∷ (ws ++ _ ∷ _) <⟨ p ⟩ _ ∷ _ ∎ drop-mid′ (swap y z p) (y ∷ z ∷ ws) (z ∷ y ∷ xs) refl refl = swap y z (drop-mid′ p _ _ refl refl) From 1d69be6c0cc1debfa0f31c40977229a335154922 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 15 Mar 2024 10:20:35 +0000 Subject: [PATCH 47/65] remove dependency on well-founded induction --- src/Data/List/Relation/Binary/Permutation/Homogeneous.agda | 1 - 1 file changed, 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 76096dba56..81685575fe 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -17,7 +17,6 @@ import Data.List.Relation.Unary.Any.Properties as Any open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) open import Data.Nat.Base using (ℕ; suc; _+_; _<_) -open import Data.Nat.Induction open import Data.Nat.Properties open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) open import Function.Base using (_∘_) From be52c13c5aefc4ec7c294f21924f4e110dd93d37 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 22 Mar 2024 13:21:50 +0000 Subject: [PATCH 48/65] further refactoring warm-ups for #1354 --- .../Relation/Binary/BagAndSetEquality.agda | 4 +- .../Binary/Permutation/Homogeneous.agda | 596 ++++++++++-------- .../Relation/Binary/Permutation/Setoid.agda | 2 +- .../Binary/Permutation/Setoid/Properties.agda | 22 +- 4 files changed, 339 insertions(+), 285 deletions(-) diff --git a/src/Data/List/Relation/Binary/BagAndSetEquality.agda b/src/Data/List/Relation/Binary/BagAndSetEquality.agda index 06f4df035a..40964d151e 100644 --- a/src/Data/List/Relation/Binary/BagAndSetEquality.agda +++ b/src/Data/List/Relation/Binary/BagAndSetEquality.agda @@ -578,10 +578,10 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys = from = ∈-resp-↭ ∘ ↭-sym from∘to : (p : xs ↭ ys) (ix : v ∈ xs) → from p (to p ix) ≡ ix - from∘to p _ = ∈-resp-↭-sym p refl + from∘to p _ = ∈-resp-↭-sym p to∘from : (p : xs ↭ ys) (iy : v ∈ ys) → to p (from p iy) ≡ iy - to∘from p _ = ∈-resp-↭-sym⁻¹ p refl + to∘from p _ = ∈-resp-↭-sym⁻¹ p ∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A} ∼bag⇒↭ {A = A} {[]} eq with refl ← empty-unique (↔-sym eq) = ↭-refl diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 81685575fe..63db71fded 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -19,7 +19,7 @@ open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) open import Data.Nat.Base using (ℕ; suc; _+_; _<_) open import Data.Nat.Properties open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) -open import Function.Base using (_∘_) +open import Function.Base using (_∘_; flip) open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_) open import Relation.Binary.Bundles using (Setoid) @@ -55,23 +55,22 @@ module _ {A : Set a} (R : Rel A r) where trans : Transitive Permutation ------------------------------------------------------------------------ --- Functions over permutations - -module _ {R : Rel A r} where +-- Map - steps : Permutation R xs ys → ℕ - steps (refl _) = 1 - steps (prep _ xs↭ys) = suc (steps xs↭ys) - steps (swap _ _ xs↭ys) = suc (steps xs↭ys) - steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs +module _ {R : Rel A r} {S : Rel A s} where --- Constructor alias + map : (R ⇒ S) → (Permutation R ⇒ Permutation S) + map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys) + map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys) + map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys) + map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) - ↭-pointwise : (Pointwise R) ⇒ Permutation R - ↭-pointwise = refl +------------------------------------------------------------------------ -- Smart inversions +module _ {R : Rel A r} where + ↭-[]-invˡ : Permutation R [] xs → xs ≡ [] ↭-[]-invˡ (refl []) = ≡.refl ↭-[]-invˡ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invˡ xs↭ys = ↭-[]-invˡ ys↭zs @@ -80,63 +79,215 @@ module _ {R : Rel A r} where ↭-[]-invʳ (refl []) = ≡.refl ↭-[]-invʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ↭-[]-invʳ xs↭ys - ¬x∷xs↭[] : ¬ (Permutation R (x ∷ xs) []) - ¬x∷xs↭[] (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ¬x∷xs↭[] xs↭ys + ¬x∷xs↭[]ˡ : ¬ (Permutation R [] (x ∷ xs)) + ¬x∷xs↭[]ˡ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invˡ xs↭ys = ¬x∷xs↭[]ˡ ys↭zs + + ¬x∷xs↭[]ʳ : ¬ (Permutation R (x ∷ xs) []) + ¬x∷xs↭[]ʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ¬x∷xs↭[]ʳ xs↭ys + + module _ (R-trans : Transitive R) where + + ↭-singleton-invˡ : Permutation R [ x ] xs → ∃ λ y → xs ≡ [ y ] × R x y + ↭-singleton-invˡ (refl (xRy ∷ [])) = _ , ≡.refl , xRy + ↭-singleton-invˡ (prep xRy p) + with ≡.refl ← ↭-[]-invˡ p = _ , ≡.refl , xRy + ↭-singleton-invˡ (trans r s) + with _ , ≡.refl , xRy ← ↭-singleton-invˡ r + with _ , ≡.refl , yRz ← ↭-singleton-invˡ s + = _ , ≡.refl , R-trans xRy yRz + + ↭-singleton-invʳ : Permutation R xs [ x ] → ∃ λ y → xs ≡ [ y ] × R y x + ↭-singleton-invʳ (refl (yRx ∷ [])) = _ , ≡.refl , yRx + ↭-singleton-invʳ (prep yRx p) + with ≡.refl ← ↭-[]-invʳ p = _ , ≡.refl , yRx + ↭-singleton-invʳ (trans r s) + with _ , ≡.refl , yRx ← ↭-singleton-invʳ s + with _ , ≡.refl , zRy ← ↭-singleton-invʳ r + = _ , ≡.refl , R-trans zRy yRx ------------------------------------------------------------------------ --- The Permutation relation is an equivalence +-- Properties of Permutation depending on suitable assumptions on R module _ {R : Rel A r} where - ↭-refl′ : Reflexive R → Reflexive (Permutation R) - ↭-refl′ R-refl = ↭-pointwise (Pointwise.refl R-refl) + private + _≋_ _↭_ : Rel (List A) _ + _≋_ = Pointwise R + _↭_ = Permutation R - sym : Symmetric R → Symmetric (Permutation R) - sym R-sym (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys) - sym R-sym (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (sym R-sym xs↭ys) - sym R-sym (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (sym R-sym xs↭ys) - sym R-sym (trans xs↭ys ys↭zs) = trans (sym R-sym ys↭zs) (sym R-sym xs↭ys) +-- Constructor alias - ↭-sym′ : Symmetric R → Symmetric (Permutation R) - ↭-sym′ = sym + ↭-pointwise : _≋_ ⇒ _↭_ + ↭-pointwise = refl - isEquivalence : Reflexive R → Symmetric R → IsEquivalence (Permutation R) - isEquivalence R-refl R-sym = record - { refl = ↭-refl′ R-refl - ; sym = ↭-sym′ R-sym - ; trans = trans - } +-- Reflexivity and its consequences - setoid : Reflexive R → Symmetric R → Setoid _ _ - setoid R-refl R-sym = record - { isEquivalence = isEquivalence R-refl R-sym - } + module _ (R-refl : Reflexive R) where + + ↭-refl′ : Reflexive _↭_ + ↭-refl′ = ↭-pointwise (Pointwise.refl R-refl) + + ↭-prep : ∀ {x} {xs ys} → Permutation R xs ys → Permutation R (x ∷ xs) (x ∷ ys) + ↭-prep xs↭ys = prep R-refl xs↭ys + + ↭-swap : ∀ x y {xs ys} → Permutation R xs ys → Permutation R (x ∷ y ∷ xs) (y ∷ x ∷ ys) + ↭-swap _ _ xs↭ys = swap R-refl R-refl xs↭ys + + ↭-shift : ∀ {v w} → R v w → ∀ xs {ys zs} → Permutation R ys zs → + Permutation R (xs List.++ v ∷ ys) (w ∷ xs List.++ zs) + ↭-shift {v} {w} v≈w [] rel = prep v≈w rel + ↭-shift {v} {w} v≈w (x ∷ xs) rel + = trans (↭-prep (↭-shift v≈w xs rel)) (↭-swap x w ↭-refl′) + + shift : ∀ {v w} → R v w → ∀ xs {ys} → + Permutation R (xs List.++ v ∷ ys) (w ∷ xs List.++ ys) + shift v≈w xs = ↭-shift v≈w xs ↭-refl′ + +-- Symmetry and its consequences + + module _ (R-sym : Symmetric R) where + + ↭-sym′ : Symmetric _↭_ + ↭-sym′ (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys) + ↭-sym′ (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (↭-sym′ xs↭ys) + ↭-sym′ (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (↭-sym′ xs↭ys) + ↭-sym′ (trans xs↭ys ys↭zs) = trans (↭-sym′ ys↭zs) (↭-sym′ xs↭ys) + +-- Transitivity and its consequences ------------------------------------------------------------------------- -- A smart version of trans that pushes `refl`s to the leaves (see #1113). -module _ {R : Rel A r} - (↭-transˡ-≋ : LeftTrans (Pointwise R) (Permutation R)) - (↭-transʳ-≋ : RightTrans (Permutation R) (Pointwise R)) - where + module _ (↭-transˡ-≋ : LeftTrans _≋_ _↭_) (↭-transʳ-≋ : RightTrans _↭_ _≋_) - ↭-trans′ : Transitive (Permutation R) - ↭-trans′ (refl xs≋ys) ys↭zs = ↭-transˡ-≋ xs≋ys ys↭zs - ↭-trans′ xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs - ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs + where + + ↭-trans′ : Transitive _↭_ + ↭-trans′ (refl xs≋ys) ys↭zs = ↭-transˡ-≋ xs≋ys ys↭zs + ↭-trans′ xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs + ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs + +-- But Left and Right Transitivity hold! + + module _ (R-trans : Transitive R) where + + private + ≋-trans : Transitive _≋_ + ≋-trans = Pointwise.transitive R-trans + + ↭-transˡ-≋ : LeftTrans _≋_ _↭_ + ↭-transˡ-≋ xs≋ys (refl ys≋zs) + = refl (≋-trans xs≋ys ys≋zs) + ↭-transˡ-≋ (x≈y ∷ xs≋ys) (prep y≈z ys↭zs) + = prep (R-trans x≈y y≈z) (↭-transˡ-≋ xs≋ys ys↭zs) + ↭-transˡ-≋ (x≈y ∷ w≈z ∷ xs≋ys) (swap eq₁ eq₂ ys↭zs) + = swap (R-trans x≈y eq₁) (R-trans w≈z eq₂) (↭-transˡ-≋ xs≋ys ys↭zs) + ↭-transˡ-≋ xs≋ys (trans ys↭ws ws↭zs) + = trans (↭-transˡ-≋ xs≋ys ys↭ws) ws↭zs + + ↭-transʳ-≋ : RightTrans _↭_ _≋_ + ↭-transʳ-≋ (refl xs≋ys) ys≋zs + = refl (≋-trans xs≋ys ys≋zs) + ↭-transʳ-≋ (prep x≈y xs↭ys) (y≈z ∷ ys≋zs) + = prep (R-trans x≈y y≈z) (↭-transʳ-≋ xs↭ys ys≋zs) + ↭-transʳ-≋ (swap eq₁ eq₂ xs↭ys) (x≈w ∷ y≈z ∷ ys≋zs) + = swap (R-trans eq₁ y≈z) (R-trans eq₂ x≈w) (↭-transʳ-≋ xs↭ys ys≋zs) + ↭-transʳ-≋ (trans xs↭ws ws↭ys) ys≋zs + = trans xs↭ws (↭-transʳ-≋ ws↭ys ys≋zs) + +-- Transitivity proper + + ↭-trans : Transitive _↭_ + ↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ ------------------------------------------------------------------------ --- Map +-- An important inversion property of Permutation R, which +-- no longer requires `steps` or well-founded induction... +------------------------------------------------------------------------ -module _ {R : Rel A r} {S : Rel A s} where +module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) + where + + private + ≋-refl : Reflexive (Pointwise R) + ≋-refl = Pointwise.refl R-refl + ↭-refl : Reflexive (Permutation R) + ↭-refl = ↭-refl′ R-refl + ≋-trans : Transitive (Pointwise R) + ≋-trans = Pointwise.transitive R-trans + _++[_]++_ = λ xs (z : A) ys → xs List.++ List.[ z ] List.++ ys + + ↭-split : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → + ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) + × Permutation R (ps List.++ qs) (as List.++ bs) + ↭-split v as bs p = helper as bs p ≋-refl + where + helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → + Pointwise R ys (as ++[ v ]++ bs) → + ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) + × Permutation R (ps List.++ qs) (as List.++ bs) + helper [] _ (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) + = [] , _ , R-trans x≈v v≈y ∷ ≋-refl , refl (≋-trans xs≋vs vs≋ys) + helper (a ∷ as) bs (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) + = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl + helper [] bs (prep {xs = xs} x≈v xs↭vs) (v≈y ∷ vs≋ys) + = [] , xs , R-trans x≈v v≈y ∷ ≋-refl , ↭-transʳ-≋ R-trans xs↭vs vs≋ys + helper (a ∷ as) bs (prep x≈v as↭vs) (v≈y ∷ vs≋ys) + with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys + = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ + helper [] [] (swap _ _ _) (_ ∷ ()) + helper [] (b ∷ bs) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + = List.[ b ] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl + , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) + helper (a ∷ []) ys (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl + , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) + helper (a ∷ b ∷ as) ys (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) + with ps , qs , eq , ↭ ← helper as ys as↭vs vs≋ys + = b ∷ a ∷ ps , qs , R-trans x≈v v≈b ∷ R-trans y≈w w≈a ∷ eq + , ↭-swap R-refl _ _ ↭ + helper as ys (trans xs↭ys ys↭zs) zs≋as++[v]++ys + with ps , qs , eq , ↭ ← helper as ys ys↭zs zs≋as++[v]++ys + with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq + = ps′ , qs′ , eq′ , ↭-trans R-trans ↭′ ↭ + + +------------------------------------------------------------------------ +-- Permutation is an equivalence satisfying another inversion property + +module _ {R : Rel A r} (R-equiv : IsEquivalence R) where + + private module ≈ = IsEquivalence R-equiv + + isEquivalence : IsEquivalence (Permutation R) + isEquivalence = record + { refl = ↭-refl′ ≈.refl + ; sym = ↭-sym′ ≈.sym + ; trans = ↭-trans ≈.trans + } + + setoid : Setoid _ _ + setoid = record { isEquivalence = isEquivalence } + + + dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → + Pointwise R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → + Permutation R (ws List.++ ys) (xs List.++ zs) + dropMiddleElement-≋ [] [] (_ ∷ eq) = ↭-pointwise eq + dropMiddleElement-≋ [] (x ∷ xs) (w≈v ∷ eq) + = ↭-transˡ-≋ ≈.trans eq (shift ≈.refl w≈v xs) + dropMiddleElement-≋ (w ∷ ws) [] (w≈x ∷ eq) + = ↭-transʳ-≋ ≈.trans (↭-sym′ ≈.sym (shift ≈.refl (≈.sym w≈x) ws)) eq + dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) = prep w≈x (dropMiddleElement-≋ ws xs eq) + + dropMiddleElement : ∀ {v : A} ws xs {ys zs} → + Permutation R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → + Permutation R (ws List.++ ys) (xs List.++ zs) + dropMiddleElement {v} ws xs {ys} {zs} p + with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans v xs zs p + = ↭-trans ≈.trans (dropMiddleElement-≋ ws ps eq) ↭ - map : (R ⇒ S) → (Permutation R ⇒ Permutation S) - map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys) - map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys) - map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys) - map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) ------------------------------------------------------------------------ -- Relationships to other predicates @@ -151,38 +302,32 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where All-resp-↭ (trans p₁ p₂) pxs = All-resp-↭ p₂ (All-resp-↭ p₁ pxs) Any-resp-↭ : (Any P) Respects (Permutation R) - Any-resp-↭ (refl xs≋ys) pxs = Pointwise.Any-resp-Pointwise resp xs≋ys pxs - Any-resp-↭ (prep x≈y p) (here px) = here (resp x≈y px) - Any-resp-↭ (prep x≈y p) (there pxs) = there (Any-resp-↭ p pxs) - Any-resp-↭ (swap x y p) (here px) = there (here (resp x px)) - Any-resp-↭ (swap x y p) (there (here px)) = here (resp y px) - Any-resp-↭ (swap x y p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) - Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) + Any-resp-↭ (refl xs≋ys) pxs = Pointwise.Any-resp-Pointwise resp xs≋ys pxs + Any-resp-↭ (prep x≈y p) (here px) = here (resp x≈y px) + Any-resp-↭ (prep x≈y p) (there pxs) = there (Any-resp-↭ p pxs) + Any-resp-↭ (swap ≈₁ ≈₂ p) (here px) = there (here (resp ≈₁ px)) + Any-resp-↭ (swap ≈₁ ≈₂ p) (there (here px)) = here (resp ≈₂ px) + Any-resp-↭ (swap ≈₁ ≈₂ p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) + Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) ------------------------------------------------------------------------ -- Two higher-dimensional properties useful in the `Propositional` case, -- specifically in the equivalence proof between `Bag` equality and `_↭_` -module _ {_≈_ : Rel A r} (isPartialEquivalence : IsPartialEquivalence _≈_) where +module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where - open IsPartialEquivalence isPartialEquivalence - renaming (sym to ≈-sym; trans to ≈-trans) + private + ∈-resp : ∀ {x} → (_≈_ x) Respects _≈_ + ∈-resp = flip ≈-trans ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) - ∈-resp-Pointwise (x≈y ∷ xs) (here ix) = here (≈-trans ix x≈y) - ∈-resp-Pointwise (x≈y ∷ xs) (there ixs) = there (∈-resp-Pointwise xs ixs) + ∈-resp-Pointwise = Pointwise.Any-resp-Pointwise ∈-resp ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) - ∈-resp-↭ (refl xs≋ys) ixs = ∈-resp-Pointwise xs≋ys ixs - ∈-resp-↭ (prep x≈y p) (here ix) = here (≈-trans ix x≈y) - ∈-resp-↭ (prep x≈y p) (there ixs) = there (∈-resp-↭ p ixs) - ∈-resp-↭ (swap x y p) (here ix) = there (here (≈-trans ix x)) - ∈-resp-↭ (swap x y p) (there (here ix)) = here (≈-trans ix y) - ∈-resp-↭ (swap x y p) (there (there ixs)) = there (there (∈-resp-↭ p ixs)) - ∈-resp-↭ (trans p₁ p₂) ixs = ∈-resp-↭ p₂ (∈-resp-↭ p₁ ixs) - - module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) + ∈-resp-↭ = Any-resp-↭ ∈-resp + module _ (≈-sym : Symmetric _≈_) + (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) where private @@ -205,43 +350,44 @@ module _ {_≈_ : Rel A r} (isPartialEquivalence : IsPartialEquivalence _≈_) w module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → ≈-trans (≈-trans p q) (≈-sym q) ≡ p) - where - ∈-resp-Pointwise-sym : (p : Pointwise _≈_ ys xs) → - {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-Pointwise p iy → - ∈-resp-Pointwise (≋-sym p) ix ≡ iy - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} {here iy} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym ix x≈y) - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} {there iys} eq - with ≡.refl ← eq = cong there (∈-resp-Pointwise-sym xs≋ys ≡.refl) + ∈-resp-Pointwise-sym : (p : Pointwise _≈_ xs ys) {ix : Any (x ≈_) xs} → + ∈-resp-Pointwise (≋-sym p) (∈-resp-Pointwise p ix) ≡ ix + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} + = cong here (≈-trans-trans-sym ix x≈y) + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} + = cong there (∈-resp-Pointwise-sym xs≋ys) - ∈-resp-↭-sym : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ∈-resp-↭-sym′ : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy - ∈-resp-↭-sym (refl xs≋ys) eq = ∈-resp-Pointwise-sym xs≋ys eq - ∈-resp-↭-sym (prep eq₁ p) {here iy} {here ix} eq + ∈-resp-↭-sym′ (refl xs≋ys) ≡.refl = ∈-resp-Pointwise-sym xs≋ys + ∈-resp-↭-sym′ (prep eq₁ p) {here iy} {here ix} eq with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) - ∈-resp-↭-sym (prep eq₁ p) {there iys} {there ixs} eq - with ≡.refl ← eq = cong there (∈-resp-↭-sym p ≡.refl) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {here iy} () - ∈-resp-↭-sym (swap eq₁ eq₂ p) {here ix} {there iys} eq + ∈-resp-↭-sym′ (prep eq₁ p) {there iys} {there ixs} eq + with ≡.refl ← eq = cong there (∈-resp-↭-sym′ p ≡.refl) + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {here iy} () + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {there iys} eq with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) - ∈-resp-↭-sym (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq - with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym p ≡.refl) - ∈-resp-↭-sym (trans p₁ p₂) eq = ∈-resp-↭-sym p₁ (∈-resp-↭-sym p₂ eq) + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq + with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym′ p ≡.refl) + ∈-resp-↭-sym′ (trans p₁ p₂) eq = ∈-resp-↭-sym′ p₁ (∈-resp-↭-sym′ p₂ eq) + + ∈-resp-↭-sym : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} → + ∈-resp-↭ (↭-sym p) (∈-resp-↭ p ix) ≡ ix + ∈-resp-↭-sym p = ∈-resp-↭-sym′ p ≡.refl - ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} {iy : Any (x ≈_) ys} → - ix ≡ ∈-resp-↭ (↭-sym p) iy → ∈-resp-↭ p ix ≡ iy - ∈-resp-↭-sym⁻¹ p eq - with eq′ ← ∈-resp-↭-sym (↭-sym p) rewrite ↭-sym-involutive′ p = eq′ eq + ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {iy : Any (x ≈_) ys} → + ∈-resp-↭ p (∈-resp-↭ (↭-sym p) iy) ≡ iy + ∈-resp-↭-sym⁻¹ p with eq′ ← ∈-resp-↭-sym′ (↭-sym p) + rewrite ↭-sym-involutive′ p = eq′ ≡.refl ------------------------------------------------------------------------ --- +-- AllPairs module _ {R : Rel A r} {S : Rel A s} (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where @@ -257,180 +403,11 @@ module _ {R : Rel A r} {S : Rel A s} AllPairs-resp-↭ (trans p₁ p₂) pxs = AllPairs-resp-↭ p₂ (AllPairs-resp-↭ p₁ pxs) ------------------------------------------------------------------------- --- Properties of steps, and properties of Permutation, --- which may depend on properties of the underlying relation ------------------------------------------------------------------------- - -module _ {R : Rel A r} where - - 0 Date: Fri, 22 Mar 2024 13:43:24 +0000 Subject: [PATCH 49/65] knock-on --- .../Binary/Permutation/Homogeneous.agda | 36 +++++++++---------- .../Permutation/Propositional/Properties.agda | 12 +++---- .../Binary/Permutation/Setoid/Properties.agda | 18 +++++----- 3 files changed, 33 insertions(+), 33 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 63db71fded..1469d1d305 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -121,6 +121,14 @@ module _ {R : Rel A r} where ↭-pointwise : _≋_ ⇒ _↭_ ↭-pointwise = refl +-- Steps: legacy definition + + steps : Permutation R xs ys → ℕ + steps (refl _) = 1 + steps (prep _ xs↭ys) = suc (steps xs↭ys) + steps (swap _ _ xs↭ys) = suc (steps xs↭ys) + steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs + -- Reflexivity and its consequences module _ (R-refl : Reflexive R) where @@ -227,30 +235,30 @@ module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) Pointwise R ys (as ++[ v ]++ bs) → ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) × Permutation R (ps List.++ qs) (as List.++ bs) + helper as bs (trans xs↭ys ys↭zs) zs≋as++[v]++ys + with ps , qs , eq , ↭ ← helper as bs ys↭zs zs≋as++[v]++ys + with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq + = ps′ , qs′ , eq′ , ↭-trans R-trans ↭′ ↭ helper [] _ (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) = [] , _ , R-trans x≈v v≈y ∷ ≋-refl , refl (≋-trans xs≋vs vs≋ys) helper (a ∷ as) bs (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) - = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl + = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl′ R-refl helper [] bs (prep {xs = xs} x≈v xs↭vs) (v≈y ∷ vs≋ys) = [] , xs , R-trans x≈v v≈y ∷ ≋-refl , ↭-transʳ-≋ R-trans xs↭vs vs≋ys helper (a ∷ as) bs (prep x≈v as↭vs) (v≈y ∷ vs≋ys) with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ - helper [] [] (swap _ _ _) (_ ∷ ()) - helper [] (b ∷ bs) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + helper [] [] (swap _ _ _) (_ ∷ ()) + helper [] (b ∷ bs) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) = List.[ b ] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) - helper (a ∷ []) ys (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + helper (a ∷ []) bs (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) - helper (a ∷ b ∷ as) ys (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) - with ps , qs , eq , ↭ ← helper as ys as↭vs vs≋ys + helper (a ∷ b ∷ as) bs (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) + with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys = b ∷ a ∷ ps , qs , R-trans x≈v v≈b ∷ R-trans y≈w w≈a ∷ eq , ↭-swap R-refl _ _ ↭ - helper as ys (trans xs↭ys ys↭zs) zs≋as++[v]++ys - with ps , qs , eq , ↭ ← helper as ys ys↭zs zs≋as++[v]++ys - with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq - = ps′ , qs′ , eq′ , ↭-trans R-trans ↭′ ↭ ------------------------------------------------------------------------ @@ -515,14 +523,6 @@ module _ (commutativeMonoid : CommutativeMonoid a r) where module Steps {R : Rel A r} where --- Function over permutations - - steps : Permutation R xs ys → ℕ - steps (refl _) = 1 - steps (prep _ xs↭ys) = suc (steps xs↭ys) - steps (swap _ _ xs↭ys) = suc (steps xs↭ys) - steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs - -- Basic property 0 Date: Fri, 22 Mar 2024 17:44:50 +0000 Subject: [PATCH 50/65] tidying up ahead of module reorganisation --- .../Binary/Permutation/Homogeneous.agda | 24 ++++++++++++------- 1 file changed, 15 insertions(+), 9 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 1469d1d305..2197ed76d5 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -67,7 +67,7 @@ module _ {R : Rel A r} {S : Rel A s} where ------------------------------------------------------------------------ --- Smart inversions +-- Inversion principles module _ {R : Rel A r} where @@ -220,8 +220,6 @@ module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) private ≋-refl : Reflexive (Pointwise R) ≋-refl = Pointwise.refl R-refl - ↭-refl : Reflexive (Permutation R) - ↭-refl = ↭-refl′ R-refl ≋-trans : Transitive (Pointwise R) ≋-trans = Pointwise.transitive R-trans _++[_]++_ = λ xs (z : A) ys → xs List.++ List.[ z ] List.++ ys @@ -249,9 +247,9 @@ module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ helper [] [] (swap _ _ _) (_ ∷ ()) - helper [] (b ∷ bs) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) - = List.[ b ] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl - , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) + helper [] (b ∷ _) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + = b ∷ [] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl + , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) helper (a ∷ []) bs (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) @@ -289,11 +287,11 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where = ↭-transʳ-≋ ≈.trans (↭-sym′ ≈.sym (shift ≈.refl (≈.sym w≈x) ws)) eq dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) = prep w≈x (dropMiddleElement-≋ ws xs eq) - dropMiddleElement : ∀ {v : A} ws xs {ys zs} → + dropMiddleElement : ∀ {x} ws xs {ys zs} → Permutation R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → Permutation R (ws List.++ ys) (xs List.++ zs) - dropMiddleElement {v} ws xs {ys} {zs} p - with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans v xs zs p + dropMiddleElement {x} ws xs {ys} {zs} p + with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p = ↭-trans ≈.trans (dropMiddleElement-≋ ws ps eq) ↭ @@ -580,8 +578,16 @@ module Steps {R : Rel A r} where -- Version 2.1 ¬x∷xs↭[] = ¬x∷xs↭[]ʳ +{-# WARNING_ON_USAGE ¬x∷xs↭[] +"Warning: ¬x∷xs↭[] was deprecated in v2.1. +Please use ¬x∷xs↭[]ʳ instead." +#-} ↭-singleton⁻¹ : {R : Rel A r} → Transitive R → ∀ {xs x} → Permutation R xs [ x ] → ∃ λ y → xs ≡ [ y ] × R y x ↭-singleton⁻¹ = ↭-singleton-invʳ +{-# WARNING_ON_USAGE ↭-singleton⁻¹ +"Warning: ↭-singleton⁻¹ was deprecated in v2.1. +Please use ↭-singleton-invʳ instead." +#-} From 487cbce8d62ac88b879659602572f1d268ae50ce Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 22 Mar 2024 18:12:31 +0000 Subject: [PATCH 51/65] more tidying up ahead of module reorganisation --- .../Binary/Permutation/Homogeneous.agda | 23 +++++++++++-------- .../Relation/Binary/Permutation/Setoid.agda | 2 +- .../Binary/Permutation/Setoid/Properties.agda | 6 ++--- 3 files changed, 18 insertions(+), 13 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 2197ed76d5..a9bd434b18 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -121,14 +121,6 @@ module _ {R : Rel A r} where ↭-pointwise : _≋_ ⇒ _↭_ ↭-pointwise = refl --- Steps: legacy definition - - steps : Permutation R xs ys → ℕ - steps (refl _) = 1 - steps (prep _ xs↭ys) = suc (steps xs↭ys) - steps (swap _ _ xs↭ys) = suc (steps xs↭ys) - steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs - -- Reflexivity and its consequences module _ (R-refl : Reflexive R) where @@ -276,7 +268,6 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where setoid : Setoid _ _ setoid = record { isEquivalence = isEquivalence } - dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → Pointwise R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → Permutation R (ws List.++ ys) (xs List.++ zs) @@ -294,6 +285,12 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p = ↭-trans ≈.trans (dropMiddleElement-≋ ws ps eq) ↭ + syntax dropMiddleElement-≋ ws xs ws++x∷ys≋xs++x∷zs + = ws ++≋[ ws++x∷ys≋xs++x∷zs ]++ xs + + syntax dropMiddleElement ws xs ws++x∷ys↭xs++x∷zs + = ws ++↭[ ws++x∷ys↭xs++x∷zs ]++ xs + ------------------------------------------------------------------------ -- Relationships to other predicates @@ -521,6 +518,14 @@ module _ (commutativeMonoid : CommutativeMonoid a r) where module Steps {R : Rel A r} where +-- Definition + + steps : Permutation R xs ys → ℕ + steps (refl _) = 1 + steps (prep _ xs↭ys) = suc (steps xs↭ys) + steps (swap _ _ xs↭ys) = suc (steps xs↭ys) + steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs + -- Basic property 0 Date: Fri, 22 Mar 2024 18:20:23 +0000 Subject: [PATCH 52/65] more tidying up ahead of module reorganisation --- src/Data/List/Relation/Binary/Permutation/Homogeneous.agda | 4 ++-- .../Relation/Binary/Permutation/Propositional/Properties.agda | 2 +- .../List/Relation/Binary/Permutation/Setoid/Properties.agda | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index a9bd434b18..841117d7ad 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -12,10 +12,10 @@ open import Algebra.Bundles using (CommutativeMonoid) open import Data.List.Base as List using (List; []; _∷_; [_]) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) -import Data.List.Relation.Unary.Any.Properties as Any open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +import Data.List.Relation.Unary.Any.Properties as Any open import Data.Nat.Base using (ℕ; suc; _+_; _<_) open import Data.Nat.Properties open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda index 0c3ed13a39..efce2cd07b 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional/Properties.agda @@ -17,8 +17,8 @@ open import Data.List.Base as List using (List; []; _∷_; [_]; _++_) open import Data.List.Membership.Propositional open import Data.List.Membership.Propositional.Properties import Data.List.Properties as List -open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.List.Relation.Unary.All using (All; []; _∷_) +open import Data.List.Relation.Unary.Any using (Any; here; there) open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) open import Function.Base using (_∘_; _⟨_⟩_) open import Level using (Level) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index 9f31aa97c1..ad7c50279f 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -17,9 +17,9 @@ open import Data.Bool.Base using (true; false) open import Data.List.Base open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise) -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) open import Data.List.Relation.Unary.All as All using (All; []; _∷_) open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) import Data.List.Relation.Unary.Unique.Setoid as Unique import Data.List.Membership.Setoid as Membership import Data.List.Properties as List From 054a7701eeb045b4274f16013b60c1cf96f6f1ad Mon Sep 17 00:00:00 2001 From: James McKinna Date: Fri, 22 Mar 2024 18:26:18 +0000 Subject: [PATCH 53/65] more tidying up ahead of module reorganisation --- .../Binary/Permutation/Homogeneous.agda | 23 +++++++++---------- 1 file changed, 11 insertions(+), 12 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 841117d7ad..f50fe44868 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -9,7 +9,7 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where open import Algebra.Bundles using (CommutativeMonoid) -open import Data.List.Base as List using (List; []; _∷_; [_]) +open import Data.List.Base as List using (List; []; _∷_; [_]; _++_; length) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) open import Data.List.Relation.Unary.All as All using (All; []; _∷_) @@ -135,13 +135,12 @@ module _ {R : Rel A r} where ↭-swap _ _ xs↭ys = swap R-refl R-refl xs↭ys ↭-shift : ∀ {v w} → R v w → ∀ xs {ys zs} → Permutation R ys zs → - Permutation R (xs List.++ v ∷ ys) (w ∷ xs List.++ zs) + Permutation R (xs ++ v ∷ ys) (w ∷ xs ++ zs) ↭-shift {v} {w} v≈w [] rel = prep v≈w rel ↭-shift {v} {w} v≈w (x ∷ xs) rel = trans (↭-prep (↭-shift v≈w xs rel)) (↭-swap x w ↭-refl′) - shift : ∀ {v w} → R v w → ∀ xs {ys} → - Permutation R (xs List.++ v ∷ ys) (w ∷ xs List.++ ys) + shift : ∀ {v w} → R v w → ∀ xs {ys} → Permutation R (xs ++ v ∷ ys) (w ∷ xs ++ ys) shift v≈w xs = ↭-shift v≈w xs ↭-refl′ -- Symmetry and its consequences @@ -214,17 +213,17 @@ module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) ≋-refl = Pointwise.refl R-refl ≋-trans : Transitive (Pointwise R) ≋-trans = Pointwise.transitive R-trans - _++[_]++_ = λ xs (z : A) ys → xs List.++ List.[ z ] List.++ ys + _++[_]++_ = λ xs (z : A) ys → xs ++ [ z ] ++ ys ↭-split : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps List.++ qs) (as List.++ bs) + × Permutation R (ps ++ qs) (as ++ bs) ↭-split v as bs p = helper as bs p ≋-refl where helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → Pointwise R ys (as ++[ v ]++ bs) → ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps List.++ qs) (as List.++ bs) + × Permutation R (ps ++ qs) (as ++ bs) helper as bs (trans xs↭ys ys↭zs) zs≋as++[v]++ys with ps , qs , eq , ↭ ← helper as bs ys↭zs zs≋as++[v]++ys with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq @@ -269,8 +268,8 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where setoid = record { isEquivalence = isEquivalence } dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → - Pointwise R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → - Permutation R (ws List.++ ys) (xs List.++ zs) + Pointwise R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → + Permutation R (ws ++ ys) (xs ++ zs) dropMiddleElement-≋ [] [] (_ ∷ eq) = ↭-pointwise eq dropMiddleElement-≋ [] (x ∷ xs) (w≈v ∷ eq) = ↭-transˡ-≋ ≈.trans eq (shift ≈.refl w≈v xs) @@ -279,8 +278,8 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) = prep w≈x (dropMiddleElement-≋ ws xs eq) dropMiddleElement : ∀ {x} ws xs {ys zs} → - Permutation R (ws List.++ x ∷ ys) (xs List.++ x ∷ zs) → - Permutation R (ws List.++ ys) (xs List.++ zs) + Permutation R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → + Permutation R (ws ++ ys) (xs ++ zs) dropMiddleElement {x} ws xs {ys} {zs} p with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p = ↭-trans ≈.trans (dropMiddleElement-≋ ws ps eq) ↭ @@ -415,7 +414,7 @@ module _ {R : Rel A r} {S : Rel A s} module _ {R : Rel A r} where - ↭-length : Permutation R xs ys → List.length xs ≡ List.length ys + ↭-length : Permutation R xs ys → length xs ≡ length ys ↭-length (refl xs≋ys) = Pointwise.Pointwise-length xs≋ys ↭-length (prep _ xs↭ys) = cong suc (↭-length xs↭ys) ↭-length (swap _ _ xs↭ys) = cong (suc ∘ suc) (↭-length xs↭ys) From 197e483c3142331005b6f7cd4f14da350604b9d8 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 10:06:02 +0000 Subject: [PATCH 54/65] further reorganisation --- .../Binary/Permutation/Homogeneous.agda | 202 ++++++++++-------- 1 file changed, 112 insertions(+), 90 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index f50fe44868..347e25eb97 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -29,7 +29,7 @@ open import Relation.Binary.Definitions using ( Reflexive; Symmetric; Transitive; LeftTrans; RightTrans ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_ ; cong) + using (_≡_ ; cong; cong₂) open import Relation.Nullary.Decidable using (yes; no; does) open import Relation.Nullary.Negation using (¬_; contradiction) open import Relation.Unary using (Pred; Decidable) @@ -66,6 +66,17 @@ module _ {R : Rel A r} {S : Rel A s} where map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) +------------------------------------------------------------------------ +-- Steps + +module _ {R : Rel A r} where + + steps : Permutation R xs ys → ℕ + steps (refl _) = 1 + steps (prep _ xs↭ys) = suc (steps xs↭ys) + steps (swap _ _ xs↭ys) = suc (steps xs↭ys) + steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs + ------------------------------------------------------------------------ -- Inversion principles @@ -121,6 +132,7 @@ module _ {R : Rel A r} where ↭-pointwise : _≋_ ⇒ _↭_ ↭-pointwise = refl +------------------------------------------------------------------------ -- Reflexivity and its consequences module _ (R-refl : Reflexive R) where @@ -143,6 +155,10 @@ module _ {R : Rel A r} where shift : ∀ {v w} → R v w → ∀ xs {ys} → Permutation R (xs ++ v ∷ ys) (w ∷ xs ++ ys) shift v≈w xs = ↭-shift v≈w xs ↭-refl′ + shift≈ : ∀ v xs {ys} → Permutation R (xs ++ v ∷ ys) (v ∷ xs ++ ys) + shift≈ v xs = shift R-refl xs + +------------------------------------------------------------------------ -- Symmetry and its consequences module _ (R-sym : Symmetric R) where @@ -153,6 +169,7 @@ module _ {R : Rel A r} where ↭-sym′ (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (↭-sym′ xs↭ys) ↭-sym′ (trans xs↭ys ys↭zs) = trans (↭-sym′ ys↭zs) (↭-sym′ xs↭ys) +------------------------------------------------------------------------ -- Transitivity and its consequences -- A smart version of trans that pushes `refl`s to the leaves (see #1113). @@ -166,7 +183,93 @@ module _ {R : Rel A r} where ↭-trans′ xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs --- But Left and Right Transitivity hold! +------------------------------------------------------------------------ +-- Two important inversion properties of Permutation R, which +-- no longer require `steps` or well-founded induction... but +-- which follow from ↭-transˡ-≋, ↭-transʳ-≋ and properties of R +------------------------------------------------------------------------ + +-- first inversion: split on a 'middle' element + + module _ (R-refl : Reflexive R) (R-trans : Transitive R) + where + + private + ≋-refl : Reflexive (Pointwise R) + ≋-refl = Pointwise.refl R-refl + ≋-trans : Transitive (Pointwise R) + ≋-trans = Pointwise.transitive R-trans + _++[_]++_ = λ xs (z : A) ys → xs ++ [ z ] ++ ys + + helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → + Pointwise R ys (as ++[ v ]++ bs) → + ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) + × Permutation R (ps ++ qs) (as ++ bs) + helper as bs (trans xs↭ys ys↭zs) zs≋as++[v]++ys + with ps , qs , eq , ↭ ← helper as bs ys↭zs zs≋as++[v]++ys + with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq + = ps′ , qs′ , eq′ , ↭-trans′ ↭′ ↭ + helper [] _ (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) + = [] , _ , R-trans x≈v v≈y ∷ ≋-refl , refl (≋-trans xs≋vs vs≋ys) + helper (a ∷ as) bs (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) + = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl′ R-refl + helper [] bs (prep {xs = xs} x≈v xs↭vs) (v≈y ∷ vs≋ys) + = [] , xs , R-trans x≈v v≈y ∷ ≋-refl , ↭-transʳ-≋ xs↭vs vs≋ys + helper (a ∷ as) bs (prep x≈v as↭vs) (v≈y ∷ vs≋ys) + with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys + = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ + helper [] [] (swap _ _ _) (_ ∷ ()) + helper [] (b ∷ _) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + = b ∷ [] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl + , ↭-prep R-refl (↭-transʳ-≋ xs↭vs vs≋ys) + helper (a ∷ []) bs (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) + = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl + , ↭-prep R-refl (↭-transʳ-≋ xs↭vs vs≋ys) + helper (a ∷ b ∷ as) bs (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) + with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys + = b ∷ a ∷ ps , qs , R-trans x≈v v≈b ∷ R-trans y≈w w≈a ∷ eq + , ↭-swap R-refl _ _ ↭ + + ↭-split : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → + ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) + × Permutation R (ps ++ qs) (as ++ bs) + ↭-split v as bs p = helper as bs p ≋-refl + + +-- second inversion: using ↭-split, drop the middle element + + module _ (R-equiv : IsEquivalence R) where + + private module ≈ = IsEquivalence R-equiv + + dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → + Pointwise R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → + Permutation R (ws ++ ys) (xs ++ zs) + dropMiddleElement-≋ [] [] (_ ∷ eq) + = ↭-pointwise eq + dropMiddleElement-≋ [] (x ∷ xs) (w≈v ∷ eq) + = ↭-transˡ-≋ eq (shift ≈.refl w≈v xs) + dropMiddleElement-≋ (w ∷ ws) [] (w≈x ∷ eq) + = ↭-transʳ-≋ (↭-sym′ ≈.sym (shift ≈.refl (≈.sym w≈x) ws)) eq + dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) + = prep w≈x (dropMiddleElement-≋ ws xs eq) + + dropMiddleElement : ∀ {x} ws xs {ys zs} → + Permutation R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → + Permutation R (ws ++ ys) (xs ++ zs) + dropMiddleElement {x} ws xs {ys} {zs} p + with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p + = ↭-trans′ (dropMiddleElement-≋ ws ps eq) ↭ + + syntax dropMiddleElement-≋ ws xs ws++x∷ys≋xs++x∷zs + = ws ++≋[ ws++x∷ys≋xs++x∷zs ]++ xs + + syntax dropMiddleElement ws xs ws++x∷ys↭xs++x∷zs + = ws ++↭[ ws++x∷ys↭xs++x∷zs ]++ xs + + +------------------------------------------------------------------------ +-- Now: Left and Right Transitivity hold! module _ (R-trans : Transitive R) where @@ -201,57 +304,7 @@ module _ {R : Rel A r} where ------------------------------------------------------------------------ --- An important inversion property of Permutation R, which --- no longer requires `steps` or well-founded induction... ------------------------------------------------------------------------- - -module _ {R : Rel A r} (R-refl : Reflexive R) (R-trans : Transitive R) - where - - private - ≋-refl : Reflexive (Pointwise R) - ≋-refl = Pointwise.refl R-refl - ≋-trans : Transitive (Pointwise R) - ≋-trans = Pointwise.transitive R-trans - _++[_]++_ = λ xs (z : A) ys → xs ++ [ z ] ++ ys - - ↭-split : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → - ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps ++ qs) (as ++ bs) - ↭-split v as bs p = helper as bs p ≋-refl - where - helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → - Pointwise R ys (as ++[ v ]++ bs) → - ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps ++ qs) (as ++ bs) - helper as bs (trans xs↭ys ys↭zs) zs≋as++[v]++ys - with ps , qs , eq , ↭ ← helper as bs ys↭zs zs≋as++[v]++ys - with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq - = ps′ , qs′ , eq′ , ↭-trans R-trans ↭′ ↭ - helper [] _ (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) - = [] , _ , R-trans x≈v v≈y ∷ ≋-refl , refl (≋-trans xs≋vs vs≋ys) - helper (a ∷ as) bs (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) - = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl′ R-refl - helper [] bs (prep {xs = xs} x≈v xs↭vs) (v≈y ∷ vs≋ys) - = [] , xs , R-trans x≈v v≈y ∷ ≋-refl , ↭-transʳ-≋ R-trans xs↭vs vs≋ys - helper (a ∷ as) bs (prep x≈v as↭vs) (v≈y ∷ vs≋ys) - with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys - = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ - helper [] [] (swap _ _ _) (_ ∷ ()) - helper [] (b ∷ _) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) - = b ∷ [] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl - , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) - helper (a ∷ []) bs (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) - = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl - , ↭-prep R-refl (↭-transʳ-≋ R-trans xs↭vs vs≋ys) - helper (a ∷ b ∷ as) bs (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) - with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys - = b ∷ a ∷ ps , qs , R-trans x≈v v≈b ∷ R-trans y≈w w≈a ∷ eq - , ↭-swap R-refl _ _ ↭ - - ------------------------------------------------------------------------- --- Permutation is an equivalence satisfying another inversion property +-- Permutation is an equivalence (Setoid version) module _ {R : Rel A r} (R-equiv : IsEquivalence R) where @@ -267,29 +320,6 @@ module _ {R : Rel A r} (R-equiv : IsEquivalence R) where setoid : Setoid _ _ setoid = record { isEquivalence = isEquivalence } - dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → - Pointwise R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → - Permutation R (ws ++ ys) (xs ++ zs) - dropMiddleElement-≋ [] [] (_ ∷ eq) = ↭-pointwise eq - dropMiddleElement-≋ [] (x ∷ xs) (w≈v ∷ eq) - = ↭-transˡ-≋ ≈.trans eq (shift ≈.refl w≈v xs) - dropMiddleElement-≋ (w ∷ ws) [] (w≈x ∷ eq) - = ↭-transʳ-≋ ≈.trans (↭-sym′ ≈.sym (shift ≈.refl (≈.sym w≈x) ws)) eq - dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) = prep w≈x (dropMiddleElement-≋ ws xs eq) - - dropMiddleElement : ∀ {x} ws xs {ys zs} → - Permutation R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → - Permutation R (ws ++ ys) (xs ++ zs) - dropMiddleElement {x} ws xs {ys} {zs} p - with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p - = ↭-trans ≈.trans (dropMiddleElement-≋ ws ps eq) ↭ - - syntax dropMiddleElement-≋ ws xs ws++x∷ys≋xs++x∷zs - = ws ++≋[ ws++x∷ys≋xs++x∷zs ]++ xs - - syntax dropMiddleElement ws xs ws++x∷ys↭xs++x∷zs - = ws ++↭[ ws++x∷ys↭xs++x∷zs ]++ xs - ------------------------------------------------------------------------ -- Relationships to other predicates @@ -319,7 +349,7 @@ module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where private - ∈-resp : ∀ {x} → (_≈_ x) Respects _≈_ + ∈-resp : ∀ {x} → (x ≈_) Respects _≈_ ∈-resp = flip ≈-trans ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) @@ -339,16 +369,16 @@ module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where ↭-sym = ↭-sym′ ≈-sym ↭-sym-involutive′ : (p : Permutation _≈_ xs ys) → ↭-sym (↭-sym p) ≡ p - ↭-sym-involutive′ (refl xs≋ys) = ≡.cong refl (≋-sym-involutive′ xs≋ys) + ↭-sym-involutive′ (refl xs≋ys) = cong refl (≋-sym-involutive′ xs≋ys) where ≋-sym-involutive′ : (p : Pointwise _≈_ xs ys) → ≋-sym (≋-sym p) ≡ p ≋-sym-involutive′ [] = ≡.refl ≋-sym-involutive′ (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y - = ≡.cong (_ ∷_) (≋-sym-involutive′ xs≋ys) - ↭-sym-involutive′ (prep eq p) = ≡.cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive′ p) + = cong (_ ∷_) (≋-sym-involutive′ xs≋ys) + ↭-sym-involutive′ (prep eq p) = cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive′ p) ↭-sym-involutive′ (swap eq₁ eq₂ p) rewrite ≈-sym-involutive eq₁ | ≈-sym-involutive eq₂ - = ≡.cong (swap _ _) (↭-sym-involutive′ p) - ↭-sym-involutive′ (trans p q) = ≡.cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) + = cong (swap _ _) (↭-sym-involutive′ p) + ↭-sym-involutive′ (trans p q) = cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → ≈-trans (≈-trans p q) (≈-sym q) ≡ p) @@ -517,14 +547,6 @@ module _ (commutativeMonoid : CommutativeMonoid a r) where module Steps {R : Rel A r} where --- Definition - - steps : Permutation R xs ys → ℕ - steps (refl _) = 1 - steps (prep _ xs↭ys) = suc (steps xs↭ys) - steps (swap _ _ xs↭ys) = suc (steps xs↭ys) - steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs - -- Basic property 0 Date: Tue, 26 Mar 2024 12:47:52 +0000 Subject: [PATCH 55/65] module restructuring --- CHANGELOG.md | 9 + .../Binary/Permutation/Homogeneous.agda | 549 +----------------- .../Permutation/Homogeneous/Properties.agda | 379 ++++++++++++ .../Homogeneous/Properties/Core.agda | 240 ++++++++ 4 files changed, 635 insertions(+), 542 deletions(-) create mode 100644 src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda create mode 100644 src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/Core.agda diff --git a/CHANGELOG.md b/CHANGELOG.md index bbc65af5d7..07a7499a30 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -30,6 +30,15 @@ Non-backwards compatible changes Other major improvements ------------------------ +* Module `Data.List.Relation.Binary.Permutation.Homogeneous` has + been comprehensively refactored, introducing + ```agda + Data.List.Relation.Binary.Permutation.Homogeneous.Properties + Data.List.Relation.Binary.Permutation.Homogeneous.Properties.Core + ``` + which then re-export to `Data.List.Relation.Binary.Permutation.Setoid.Properties` + and `Data.List.Relation.Binary.Permutation.Propositional.Properties` + Deprecated modules ------------------ diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 347e25eb97..5e8e379bdd 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -8,7 +8,6 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where -open import Algebra.Bundles using (CommutativeMonoid) open import Data.List.Base as List using (List; []; _∷_; [_]; _++_; length) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) @@ -66,535 +65,6 @@ module _ {R : Rel A r} {S : Rel A s} where map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs) ------------------------------------------------------------------------- --- Steps - -module _ {R : Rel A r} where - - steps : Permutation R xs ys → ℕ - steps (refl _) = 1 - steps (prep _ xs↭ys) = suc (steps xs↭ys) - steps (swap _ _ xs↭ys) = suc (steps xs↭ys) - steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs - ------------------------------------------------------------------------- --- Inversion principles - -module _ {R : Rel A r} where - - ↭-[]-invˡ : Permutation R [] xs → xs ≡ [] - ↭-[]-invˡ (refl []) = ≡.refl - ↭-[]-invˡ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invˡ xs↭ys = ↭-[]-invˡ ys↭zs - - ↭-[]-invʳ : Permutation R xs [] → xs ≡ [] - ↭-[]-invʳ (refl []) = ≡.refl - ↭-[]-invʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ↭-[]-invʳ xs↭ys - - ¬x∷xs↭[]ˡ : ¬ (Permutation R [] (x ∷ xs)) - ¬x∷xs↭[]ˡ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invˡ xs↭ys = ¬x∷xs↭[]ˡ ys↭zs - - ¬x∷xs↭[]ʳ : ¬ (Permutation R (x ∷ xs) []) - ¬x∷xs↭[]ʳ (trans xs↭ys ys↭zs) with ≡.refl ← ↭-[]-invʳ ys↭zs = ¬x∷xs↭[]ʳ xs↭ys - - module _ (R-trans : Transitive R) where - - ↭-singleton-invˡ : Permutation R [ x ] xs → ∃ λ y → xs ≡ [ y ] × R x y - ↭-singleton-invˡ (refl (xRy ∷ [])) = _ , ≡.refl , xRy - ↭-singleton-invˡ (prep xRy p) - with ≡.refl ← ↭-[]-invˡ p = _ , ≡.refl , xRy - ↭-singleton-invˡ (trans r s) - with _ , ≡.refl , xRy ← ↭-singleton-invˡ r - with _ , ≡.refl , yRz ← ↭-singleton-invˡ s - = _ , ≡.refl , R-trans xRy yRz - - ↭-singleton-invʳ : Permutation R xs [ x ] → ∃ λ y → xs ≡ [ y ] × R y x - ↭-singleton-invʳ (refl (yRx ∷ [])) = _ , ≡.refl , yRx - ↭-singleton-invʳ (prep yRx p) - with ≡.refl ← ↭-[]-invʳ p = _ , ≡.refl , yRx - ↭-singleton-invʳ (trans r s) - with _ , ≡.refl , yRx ← ↭-singleton-invʳ s - with _ , ≡.refl , zRy ← ↭-singleton-invʳ r - = _ , ≡.refl , R-trans zRy yRx - - ------------------------------------------------------------------------- --- Properties of Permutation depending on suitable assumptions on R - -module _ {R : Rel A r} where - - private - _≋_ _↭_ : Rel (List A) _ - _≋_ = Pointwise R - _↭_ = Permutation R - --- Constructor alias - - ↭-pointwise : _≋_ ⇒ _↭_ - ↭-pointwise = refl - ------------------------------------------------------------------------- --- Reflexivity and its consequences - - module _ (R-refl : Reflexive R) where - - ↭-refl′ : Reflexive _↭_ - ↭-refl′ = ↭-pointwise (Pointwise.refl R-refl) - - ↭-prep : ∀ {x} {xs ys} → Permutation R xs ys → Permutation R (x ∷ xs) (x ∷ ys) - ↭-prep xs↭ys = prep R-refl xs↭ys - - ↭-swap : ∀ x y {xs ys} → Permutation R xs ys → Permutation R (x ∷ y ∷ xs) (y ∷ x ∷ ys) - ↭-swap _ _ xs↭ys = swap R-refl R-refl xs↭ys - - ↭-shift : ∀ {v w} → R v w → ∀ xs {ys zs} → Permutation R ys zs → - Permutation R (xs ++ v ∷ ys) (w ∷ xs ++ zs) - ↭-shift {v} {w} v≈w [] rel = prep v≈w rel - ↭-shift {v} {w} v≈w (x ∷ xs) rel - = trans (↭-prep (↭-shift v≈w xs rel)) (↭-swap x w ↭-refl′) - - shift : ∀ {v w} → R v w → ∀ xs {ys} → Permutation R (xs ++ v ∷ ys) (w ∷ xs ++ ys) - shift v≈w xs = ↭-shift v≈w xs ↭-refl′ - - shift≈ : ∀ v xs {ys} → Permutation R (xs ++ v ∷ ys) (v ∷ xs ++ ys) - shift≈ v xs = shift R-refl xs - ------------------------------------------------------------------------- --- Symmetry and its consequences - - module _ (R-sym : Symmetric R) where - - ↭-sym′ : Symmetric _↭_ - ↭-sym′ (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys) - ↭-sym′ (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (↭-sym′ xs↭ys) - ↭-sym′ (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (↭-sym′ xs↭ys) - ↭-sym′ (trans xs↭ys ys↭zs) = trans (↭-sym′ ys↭zs) (↭-sym′ xs↭ys) - ------------------------------------------------------------------------- --- Transitivity and its consequences - --- A smart version of trans that pushes `refl`s to the leaves (see #1113). - - module _ (↭-transˡ-≋ : LeftTrans _≋_ _↭_) (↭-transʳ-≋ : RightTrans _↭_ _≋_) - - where - - ↭-trans′ : Transitive _↭_ - ↭-trans′ (refl xs≋ys) ys↭zs = ↭-transˡ-≋ xs≋ys ys↭zs - ↭-trans′ xs↭ys (refl ys≋zs) = ↭-transʳ-≋ xs↭ys ys≋zs - ↭-trans′ xs↭ys ys↭zs = trans xs↭ys ys↭zs - ------------------------------------------------------------------------- --- Two important inversion properties of Permutation R, which --- no longer require `steps` or well-founded induction... but --- which follow from ↭-transˡ-≋, ↭-transʳ-≋ and properties of R ------------------------------------------------------------------------- - --- first inversion: split on a 'middle' element - - module _ (R-refl : Reflexive R) (R-trans : Transitive R) - where - - private - ≋-refl : Reflexive (Pointwise R) - ≋-refl = Pointwise.refl R-refl - ≋-trans : Transitive (Pointwise R) - ≋-trans = Pointwise.transitive R-trans - _++[_]++_ = λ xs (z : A) ys → xs ++ [ z ] ++ ys - - helper : ∀ as bs {xs ys} (p : Permutation R xs ys) → - Pointwise R ys (as ++[ v ]++ bs) → - ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps ++ qs) (as ++ bs) - helper as bs (trans xs↭ys ys↭zs) zs≋as++[v]++ys - with ps , qs , eq , ↭ ← helper as bs ys↭zs zs≋as++[v]++ys - with ps′ , qs′ , eq′ , ↭′ ← helper ps qs xs↭ys eq - = ps′ , qs′ , eq′ , ↭-trans′ ↭′ ↭ - helper [] _ (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) - = [] , _ , R-trans x≈v v≈y ∷ ≋-refl , refl (≋-trans xs≋vs vs≋ys) - helper (a ∷ as) bs (refl (x≈v ∷ xs≋vs)) (v≈y ∷ vs≋ys) - = _ ∷ as , bs , R-trans x≈v v≈y ∷ ≋-trans xs≋vs vs≋ys , ↭-refl′ R-refl - helper [] bs (prep {xs = xs} x≈v xs↭vs) (v≈y ∷ vs≋ys) - = [] , xs , R-trans x≈v v≈y ∷ ≋-refl , ↭-transʳ-≋ xs↭vs vs≋ys - helper (a ∷ as) bs (prep x≈v as↭vs) (v≈y ∷ vs≋ys) - with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys - = a ∷ ps , qs , R-trans x≈v v≈y ∷ eq , prep R-refl ↭ - helper [] [] (swap _ _ _) (_ ∷ ()) - helper [] (b ∷ _) (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) - = b ∷ [] , _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl - , ↭-prep R-refl (↭-transʳ-≋ xs↭vs vs≋ys) - helper (a ∷ []) bs (swap x≈v y≈w xs↭vs) (w≈z ∷ v≈y ∷ vs≋ys) - = [] , a ∷ _ , R-trans x≈v v≈y ∷ R-trans y≈w w≈z ∷ ≋-refl - , ↭-prep R-refl (↭-transʳ-≋ xs↭vs vs≋ys) - helper (a ∷ b ∷ as) bs (swap x≈v y≈w as↭vs) (w≈a ∷ v≈b ∷ vs≋ys) - with ps , qs , eq , ↭ ← helper as bs as↭vs vs≋ys - = b ∷ a ∷ ps , qs , R-trans x≈v v≈b ∷ R-trans y≈w w≈a ∷ eq - , ↭-swap R-refl _ _ ↭ - - ↭-split : ∀ v as bs {xs} → Permutation R xs (as ++[ v ]++ bs) → - ∃₂ λ ps qs → Pointwise R xs (ps ++[ v ]++ qs) - × Permutation R (ps ++ qs) (as ++ bs) - ↭-split v as bs p = helper as bs p ≋-refl - - --- second inversion: using ↭-split, drop the middle element - - module _ (R-equiv : IsEquivalence R) where - - private module ≈ = IsEquivalence R-equiv - - dropMiddleElement-≋ : ∀ {x} ws xs {ys} {zs} → - Pointwise R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → - Permutation R (ws ++ ys) (xs ++ zs) - dropMiddleElement-≋ [] [] (_ ∷ eq) - = ↭-pointwise eq - dropMiddleElement-≋ [] (x ∷ xs) (w≈v ∷ eq) - = ↭-transˡ-≋ eq (shift ≈.refl w≈v xs) - dropMiddleElement-≋ (w ∷ ws) [] (w≈x ∷ eq) - = ↭-transʳ-≋ (↭-sym′ ≈.sym (shift ≈.refl (≈.sym w≈x) ws)) eq - dropMiddleElement-≋ (w ∷ ws) (x ∷ xs) (w≈x ∷ eq) - = prep w≈x (dropMiddleElement-≋ ws xs eq) - - dropMiddleElement : ∀ {x} ws xs {ys zs} → - Permutation R (ws ++ x ∷ ys) (xs ++ x ∷ zs) → - Permutation R (ws ++ ys) (xs ++ zs) - dropMiddleElement {x} ws xs {ys} {zs} p - with ps , qs , eq , ↭ ← ↭-split ≈.refl ≈.trans x xs zs p - = ↭-trans′ (dropMiddleElement-≋ ws ps eq) ↭ - - syntax dropMiddleElement-≋ ws xs ws++x∷ys≋xs++x∷zs - = ws ++≋[ ws++x∷ys≋xs++x∷zs ]++ xs - - syntax dropMiddleElement ws xs ws++x∷ys↭xs++x∷zs - = ws ++↭[ ws++x∷ys↭xs++x∷zs ]++ xs - - ------------------------------------------------------------------------- --- Now: Left and Right Transitivity hold! - - module _ (R-trans : Transitive R) where - - private - ≋-trans : Transitive _≋_ - ≋-trans = Pointwise.transitive R-trans - - ↭-transˡ-≋ : LeftTrans _≋_ _↭_ - ↭-transˡ-≋ xs≋ys (refl ys≋zs) - = refl (≋-trans xs≋ys ys≋zs) - ↭-transˡ-≋ (x≈y ∷ xs≋ys) (prep y≈z ys↭zs) - = prep (R-trans x≈y y≈z) (↭-transˡ-≋ xs≋ys ys↭zs) - ↭-transˡ-≋ (x≈y ∷ w≈z ∷ xs≋ys) (swap eq₁ eq₂ ys↭zs) - = swap (R-trans x≈y eq₁) (R-trans w≈z eq₂) (↭-transˡ-≋ xs≋ys ys↭zs) - ↭-transˡ-≋ xs≋ys (trans ys↭ws ws↭zs) - = trans (↭-transˡ-≋ xs≋ys ys↭ws) ws↭zs - - ↭-transʳ-≋ : RightTrans _↭_ _≋_ - ↭-transʳ-≋ (refl xs≋ys) ys≋zs - = refl (≋-trans xs≋ys ys≋zs) - ↭-transʳ-≋ (prep x≈y xs↭ys) (y≈z ∷ ys≋zs) - = prep (R-trans x≈y y≈z) (↭-transʳ-≋ xs↭ys ys≋zs) - ↭-transʳ-≋ (swap eq₁ eq₂ xs↭ys) (x≈w ∷ y≈z ∷ ys≋zs) - = swap (R-trans eq₁ y≈z) (R-trans eq₂ x≈w) (↭-transʳ-≋ xs↭ys ys≋zs) - ↭-transʳ-≋ (trans xs↭ws ws↭ys) ys≋zs - = trans xs↭ws (↭-transʳ-≋ ws↭ys ys≋zs) - --- Transitivity proper - - ↭-trans : Transitive _↭_ - ↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ - - ------------------------------------------------------------------------- --- Permutation is an equivalence (Setoid version) - -module _ {R : Rel A r} (R-equiv : IsEquivalence R) where - - private module ≈ = IsEquivalence R-equiv - - isEquivalence : IsEquivalence (Permutation R) - isEquivalence = record - { refl = ↭-refl′ ≈.refl - ; sym = ↭-sym′ ≈.sym - ; trans = ↭-trans ≈.trans - } - - setoid : Setoid _ _ - setoid = record { isEquivalence = isEquivalence } - - ------------------------------------------------------------------------- --- Relationships to other predicates ------------------------------------------------------------------------- - -module _ {R : Rel A r} {P : Pred A p} (resp : P Respects R) where - - All-resp-↭ : (All P) Respects (Permutation R) - All-resp-↭ (refl xs≋ys) pxs = Pointwise.All-resp-Pointwise resp xs≋ys pxs - All-resp-↭ (prep x≈y p) (px ∷ pxs) = resp x≈y px ∷ All-resp-↭ p pxs - All-resp-↭ (swap ≈₁ ≈₂ p) (px ∷ py ∷ pxs) = resp ≈₂ py ∷ resp ≈₁ px ∷ All-resp-↭ p pxs - All-resp-↭ (trans p₁ p₂) pxs = All-resp-↭ p₂ (All-resp-↭ p₁ pxs) - - Any-resp-↭ : (Any P) Respects (Permutation R) - Any-resp-↭ (refl xs≋ys) pxs = Pointwise.Any-resp-Pointwise resp xs≋ys pxs - Any-resp-↭ (prep x≈y p) (here px) = here (resp x≈y px) - Any-resp-↭ (prep x≈y p) (there pxs) = there (Any-resp-↭ p pxs) - Any-resp-↭ (swap ≈₁ ≈₂ p) (here px) = there (here (resp ≈₁ px)) - Any-resp-↭ (swap ≈₁ ≈₂ p) (there (here px)) = here (resp ≈₂ px) - Any-resp-↭ (swap ≈₁ ≈₂ p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) - Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) - ------------------------------------------------------------------------- --- Two higher-dimensional properties useful in the `Propositional` case, --- specifically in the equivalence proof between `Bag` equality and `_↭_` - -module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where - - private - ∈-resp : ∀ {x} → (x ≈_) Respects _≈_ - ∈-resp = flip ≈-trans - - ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) - ∈-resp-Pointwise = Pointwise.Any-resp-Pointwise ∈-resp - - ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) - ∈-resp-↭ = Any-resp-↭ ∈-resp - - module _ (≈-sym : Symmetric _≈_) - (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) - where - - private - ≋-sym : Symmetric (Pointwise _≈_) - ≋-sym = Pointwise.symmetric ≈-sym - ↭-sym : Symmetric (Permutation _≈_) - ↭-sym = ↭-sym′ ≈-sym - - ↭-sym-involutive′ : (p : Permutation _≈_ xs ys) → ↭-sym (↭-sym p) ≡ p - ↭-sym-involutive′ (refl xs≋ys) = cong refl (≋-sym-involutive′ xs≋ys) - where - ≋-sym-involutive′ : (p : Pointwise _≈_ xs ys) → ≋-sym (≋-sym p) ≡ p - ≋-sym-involutive′ [] = ≡.refl - ≋-sym-involutive′ (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y - = cong (_ ∷_) (≋-sym-involutive′ xs≋ys) - ↭-sym-involutive′ (prep eq p) = cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive′ p) - ↭-sym-involutive′ (swap eq₁ eq₂ p) rewrite ≈-sym-involutive eq₁ | ≈-sym-involutive eq₂ - = cong (swap _ _) (↭-sym-involutive′ p) - ↭-sym-involutive′ (trans p q) = cong₂ trans (↭-sym-involutive′ p) (↭-sym-involutive′ q) - - module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → - ≈-trans (≈-trans p q) (≈-sym q) ≡ p) - where - - ∈-resp-Pointwise-sym : (p : Pointwise _≈_ xs ys) {ix : Any (x ≈_) xs} → - ∈-resp-Pointwise (≋-sym p) (∈-resp-Pointwise p ix) ≡ ix - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} - = cong here (≈-trans-trans-sym ix x≈y) - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} - = cong there (∈-resp-Pointwise-sym xs≋ys) - - ∈-resp-↭-sym′ : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy - ∈-resp-↭-sym′ (refl xs≋ys) ≡.refl = ∈-resp-Pointwise-sym xs≋ys - ∈-resp-↭-sym′ (prep eq₁ p) {here iy} {here ix} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) - ∈-resp-↭-sym′ (prep eq₁ p) {there iys} {there ixs} eq - with ≡.refl ← eq = cong there (∈-resp-↭-sym′ p ≡.refl) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {here iy} () - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {there iys} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq - with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq - with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym′ p ≡.refl) - ∈-resp-↭-sym′ (trans p₁ p₂) eq = ∈-resp-↭-sym′ p₁ (∈-resp-↭-sym′ p₂ eq) - - ∈-resp-↭-sym : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} → - ∈-resp-↭ (↭-sym p) (∈-resp-↭ p ix) ≡ ix - ∈-resp-↭-sym p = ∈-resp-↭-sym′ p ≡.refl - - ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {iy : Any (x ≈_) ys} → - ∈-resp-↭ p (∈-resp-↭ (↭-sym p) iy) ≡ iy - ∈-resp-↭-sym⁻¹ p with eq′ ← ∈-resp-↭-sym′ (↭-sym p) - rewrite ↭-sym-involutive′ p = eq′ ≡.refl - - ------------------------------------------------------------------------- --- AllPairs - -module _ {R : Rel A r} {S : Rel A s} - (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where - - AllPairs-resp-↭ : (AllPairs S) Respects (Permutation R) - AllPairs-resp-↭ (refl xs≋ys) pxs = - Pointwise.AllPairs-resp-Pointwise resp xs≋ys pxs - AllPairs-resp-↭ (prep x≈y p) (∼ ∷ pxs) = - All-resp-↭ rʳ p (All.map (rˡ x≈y) ∼) ∷ AllPairs-resp-↭ p pxs - AllPairs-resp-↭ (swap eq₁ eq₂ p) ((∼₁ ∷ ∼₂) ∷ ∼₃ ∷ pxs) = - (sym (rʳ eq₂ (rˡ eq₁ ∼₁)) ∷ All-resp-↭ rʳ p (All.map (rˡ eq₂) ∼₃)) ∷ - All-resp-↭ rʳ p (All.map (rˡ eq₁) ∼₂) ∷ AllPairs-resp-↭ p pxs - AllPairs-resp-↭ (trans p₁ p₂) pxs = - AllPairs-resp-↭ p₂ (AllPairs-resp-↭ p₁ pxs) - - ------------------------------------------------------------------------- --- Properties of List functions, possibly depending on properties of R ------------------------------------------------------------------------- - --- length - -module _ {R : Rel A r} where - - ↭-length : Permutation R xs ys → length xs ≡ length ys - ↭-length (refl xs≋ys) = Pointwise.Pointwise-length xs≋ys - ↭-length (prep _ xs↭ys) = cong suc (↭-length xs↭ys) - ↭-length (swap _ _ xs↭ys) = cong (suc ∘ suc) (↭-length xs↭ys) - ↭-length (trans xs↭ys ys↭zs) = ≡.trans (↭-length xs↭ys) (↭-length ys↭zs) - --- map - -module _ {R : Rel A r} {S : Rel B s} {f} (pres : f Preserves R ⟶ S) where - - map⁺ : Permutation R xs ys → Permutation S (List.map f xs) (List.map f ys) - - map⁺ (refl xs≋ys) = refl (Pointwise.map⁺ _ _ (Pointwise.map pres xs≋ys)) - map⁺ (prep x xs↭ys) = prep (pres x) (map⁺ xs↭ys) - map⁺ (swap x y xs↭ys) = swap (pres x) (pres y) (map⁺ xs↭ys) - map⁺ (trans xs↭ys ys↭zs) = trans (map⁺ xs↭ys) (map⁺ ys↭zs) -{- - -- permutations preserve 'being a mapped list' - ↭-map-inv : ∀ {zs : List B} → Permutation S (List.map f xs) zs → - ∃ λ ys → zs ≡ List.map f ys × Permutation R xs ys - ↭-map-inv p = {!p!} --} - --- _++_ - -module _ {R : Rel A r} (R-refl : Reflexive R) where - - ++⁺ʳ : ∀ zs → Permutation R xs ys → - Permutation R (xs List.++ zs) (ys List.++ zs) - ++⁺ʳ zs (refl xs≋ys) = refl (Pointwise.++⁺ xs≋ys (Pointwise.refl R-refl)) - ++⁺ʳ zs (prep x xs↭ys) = prep x (++⁺ʳ zs xs↭ys) - ++⁺ʳ zs (swap x y xs↭ys) = swap x y (++⁺ʳ zs xs↭ys) - ++⁺ʳ zs (trans xs↭ys ys↭zs) = trans (++⁺ʳ zs xs↭ys) (++⁺ʳ zs ys↭zs) - - --- filter - -module _ {R : Rel A r} (R-sym : Symmetric R) - {p} {P : Pred A p} (P? : Decidable P) (P≈ : P Respects R) where - - filter⁺ : Permutation R xs ys → - Permutation R (List.filter P? xs) (List.filter P? ys) - filter⁺ (refl xs≋ys) = refl (Pointwise.filter⁺ P? P? P≈ (P≈ ∘ R-sym) xs≋ys) - filter⁺ (trans xs↭zs zs↭ys) = trans (filter⁺ xs↭zs) (filter⁺ zs↭ys) - filter⁺ {x ∷ xs} {y ∷ ys} (prep x≈y xs↭ys) with P? x | P? y - ... | yes _ | yes _ = prep x≈y (filter⁺ xs↭ys) - ... | yes Px | no ¬Py = contradiction (P≈ x≈y Px) ¬Py - ... | no ¬Px | yes Py = contradiction (P≈ (R-sym x≈y) Py) ¬Px - ... | no _ | no _ = filter⁺ xs↭ys - filter⁺ {x ∷ w ∷ xs} {y ∷ z ∷ ys} (swap x≈z w≈y xs↭ys) with P? x | P? y - filter⁺ {x ∷ w ∷ xs} {y ∷ z ∷ ys} (swap x≈z w≈y xs↭ys) | no ¬Px | no ¬Py - with P? z | P? w - ... | _ | yes Pw = contradiction (P≈ w≈y Pw) ¬Py - ... | yes Pz | _ = contradiction (P≈ (R-sym x≈z) Pz) ¬Px - ... | no _ | no _ = filter⁺ xs↭ys - filter⁺ {x ∷ w ∷ xs} {y ∷ z ∷ ys} (swap x≈z w≈y xs↭ys) | no ¬Px | yes Py - with P? z | P? w - ... | _ | no ¬Pw = contradiction (P≈ (R-sym w≈y) Py) ¬Pw - ... | yes Pz | _ = contradiction (P≈ (R-sym x≈z) Pz) ¬Px - ... | no _ | yes _ = prep w≈y (filter⁺ xs↭ys) - filter⁺ {x ∷ w ∷ xs} {y ∷ z ∷ ys} (swap x≈z w≈y xs↭ys) | yes Px | no ¬Py - with P? z | P? w - ... | no ¬Pz | _ = contradiction (P≈ x≈z Px) ¬Pz - ... | _ | yes Pw = contradiction (P≈ w≈y Pw) ¬Py - ... | yes _ | no _ = prep x≈z (filter⁺ xs↭ys) - filter⁺ {x ∷ w ∷ xs} {y ∷ z ∷ ys} (swap x≈z w≈y xs↭ys) | yes Px | yes Py - with P? z | P? w - ... | no ¬Pz | _ = contradiction (P≈ x≈z Px) ¬Pz - ... | _ | no ¬Pw = contradiction (P≈ (R-sym w≈y) Py) ¬Pw - ... | yes _ | yes _ = swap x≈z w≈y (filter⁺ xs↭ys) - - ------------------------------------------------------------------------- --- foldr of Commutative Monoid - -module _ (commutativeMonoid : CommutativeMonoid a r) where - open module CM = CommutativeMonoid commutativeMonoid - - private foldr = List.foldr - - foldr-commMonoid : Permutation _≈_ xs ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys - foldr-commMonoid (refl xs≋ys) = Pointwise.foldr⁺ ∙-cong CM.refl xs≋ys - foldr-commMonoid (prep x≈y xs↭ys) = ∙-cong x≈y (foldr-commMonoid xs↭ys) - foldr-commMonoid (swap {xs} {ys} {x} {y} {x′} {y′} x≈x′ y≈y′ xs↭ys) = begin - x ∙ (y ∙ foldr _∙_ ε xs) ≈⟨ ∙-congˡ (∙-congˡ (foldr-commMonoid xs↭ys)) ⟩ - x ∙ (y ∙ foldr _∙_ ε ys) ≈⟨ assoc x y (foldr _∙_ ε ys) ⟨ - (x ∙ y) ∙ foldr _∙_ ε ys ≈⟨ ∙-congʳ (comm x y) ⟩ - (y ∙ x) ∙ foldr _∙_ ε ys ≈⟨ ∙-congʳ (∙-cong y≈y′ x≈x′) ⟩ - (y′ ∙ x′) ∙ foldr _∙_ ε ys ≈⟨ assoc y′ x′ (foldr _∙_ ε ys) ⟩ - y′ ∙ (x′ ∙ foldr _∙_ ε ys) ∎ - where open ≈-Reasoning CM.setoid - foldr-commMonoid (trans xs↭ys ys↭zs) = - CM.trans (foldr-commMonoid xs↭ys) (foldr-commMonoid ys↭zs) - - ------------------------------------------------------------------------- --- Properties of steps, and related properties of Permutation --- previously required for proofs by well-founded induction --- rendered obsolete/deprecatable by ↭-transˡ-≋ , ↭-transʳ-≋ ------------------------------------------------------------------------- - -module Steps {R : Rel A r} where - --- Basic property - - 0 Date: Tue, 26 Mar 2024 13:03:26 +0000 Subject: [PATCH 56/65] typechecking proofs of involutive symmetry fails to terminate after module restructuring --- .../Relation/Binary/Permutation/Homogeneous/Properties.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda index f9e531c99a..d9379fb1a7 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda @@ -230,7 +230,7 @@ module _ (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where ------------------------------------------------------------------------ -- Two higher-dimensional properties useful in the `Propositional` case, -- specifically in the equivalence proof between `Bag` equality and `_↭_` - +{- module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where private @@ -298,7 +298,7 @@ module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where ∈-resp-↭-sym⁻¹ p with eq′ ← ∈-resp-↭-sym′ (↭-sym p) rewrite ↭-sym-involutive p = eq′ ≡.refl - +-} ------------------------------------------------------------------------ -- Properties of steps, and related properties of Permutation -- previously required for proofs by well-founded induction From b48e06facd71db2589bc5d73488364ba1a995440 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:12:55 +0000 Subject: [PATCH 57/65] factored out the higher-dimensional properties --- .../Permutation/Homogeneous/Properties.agda | 91 ++++--------------- .../Homogeneous/Properties/Core.agda | 14 +-- 2 files changed, 25 insertions(+), 80 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda index d9379fb1a7..4e43c50876 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda @@ -19,7 +19,7 @@ open import Data.List.Relation.Unary.Any as Any using (Any; here; there) import Data.List.Relation.Unary.Any.Properties as Any open import Data.Nat.Base using (ℕ; suc; _+_; _<_) import Data.Nat.Properties as ℕ -open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) +open import Data.Product.Base using (_,_; _×_; ∃) open import Function.Base using (_∘_; flip) open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _Preserves_⟶_) @@ -44,7 +44,7 @@ private x y z v w : A P : Pred A p R : Rel A r - S : Rel A s + S : Rel A s ------------------------------------------------------------------------ -- Re-export `Core` properties @@ -212,10 +212,24 @@ module _ (resp : P Respects R) where Any-resp-↭ (swap ≈₁ ≈₂ p) (there (there pxs)) = there (there (Any-resp-↭ p pxs)) Any-resp-↭ (trans p₁ p₂) pxs = Any-resp-↭ p₂ (Any-resp-↭ p₁ pxs) +-- Membership + +module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where + + private + ∈-resp : ∀ {x} → (x ≈_) Respects _≈_ + ∈-resp = flip ≈-trans + + ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) + ∈-resp-Pointwise = Pointwise.Any-resp-Pointwise ∈-resp + + ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) + ∈-resp-↭ = Any-resp-↭ ∈-resp + -- AllPairs module _ (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where - + AllPairs-resp-↭ : (AllPairs S) Respects (Permutation R) AllPairs-resp-↭ (refl xs≋ys) pxs = Pointwise.AllPairs-resp-Pointwise resp xs≋ys pxs @@ -227,78 +241,7 @@ module _ (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where AllPairs-resp-↭ (trans p₁ p₂) pxs = AllPairs-resp-↭ p₂ (AllPairs-resp-↭ p₁ pxs) ------------------------------------------------------------------------- --- Two higher-dimensional properties useful in the `Propositional` case, --- specifically in the equivalence proof between `Bag` equality and `_↭_` -{- -module _ {_≈_ : Rel A r} (≈-trans : Transitive _≈_) where - - private - ∈-resp : ∀ {x} → (x ≈_) Respects _≈_ - ∈-resp = flip ≈-trans - - ∈-resp-Pointwise : (Any (x ≈_)) Respects (Pointwise _≈_) - ∈-resp-Pointwise = Pointwise.Any-resp-Pointwise ∈-resp - - ∈-resp-↭ : (Any (x ≈_)) Respects (Permutation _≈_) - ∈-resp-↭ = Any-resp-↭ ∈-resp - module _ (≈-sym : Symmetric _≈_) - (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) - where - - open Symmetry ≈-sym using (≋-sym; ↭-sym) - - ↭-sym-involutive : (p : Permutation _≈_ xs ys) → ↭-sym (↭-sym p) ≡ p - ↭-sym-involutive (refl xs≋ys) = cong refl (≋-sym-involutive xs≋ys) - where - ≋-sym-involutive : (p : Pointwise _≈_ xs ys) → ≋-sym (≋-sym p) ≡ p - ≋-sym-involutive [] = ≡.refl - ≋-sym-involutive (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y - = cong (_ ∷_) (≋-sym-involutive xs≋ys) - ↭-sym-involutive (prep eq p) = cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive p) - ↭-sym-involutive (swap eq₁ eq₂ p) rewrite ≈-sym-involutive eq₁ | ≈-sym-involutive eq₂ - = cong (swap _ _) (↭-sym-involutive p) - ↭-sym-involutive (trans p q) = cong₂ trans (↭-sym-involutive p) (↭-sym-involutive q) - - module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → - ≈-trans (≈-trans p q) (≈-sym q) ≡ p) - where - - ∈-resp-Pointwise-sym : (p : Pointwise _≈_ xs ys) {ix : Any (x ≈_) xs} → - ∈-resp-Pointwise (≋-sym p) (∈-resp-Pointwise p ix) ≡ ix - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} - = cong here (≈-trans-trans-sym ix x≈y) - ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} - = cong there (∈-resp-Pointwise-sym xs≋ys) - - ∈-resp-↭-sym′ : (p : Permutation _≈_ ys xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → - ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy - ∈-resp-↭-sym′ (refl xs≋ys) ≡.refl = ∈-resp-Pointwise-sym xs≋ys - ∈-resp-↭-sym′ (prep eq₁ p) {here iy} {here ix} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) - ∈-resp-↭-sym′ (prep eq₁ p) {there iys} {there ixs} eq - with ≡.refl ← eq = cong there (∈-resp-↭-sym′ p ≡.refl) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {here iy} () - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {there iys} eq - with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq - with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) - ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq - with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym′ p ≡.refl) - ∈-resp-↭-sym′ (trans p₁ p₂) eq = ∈-resp-↭-sym′ p₁ (∈-resp-↭-sym′ p₂ eq) - - ∈-resp-↭-sym : (p : Permutation _≈_ xs ys) {ix : Any (x ≈_) xs} → - ∈-resp-↭ (↭-sym p) (∈-resp-↭ p ix) ≡ ix - ∈-resp-↭-sym p = ∈-resp-↭-sym′ p ≡.refl - - ∈-resp-↭-sym⁻¹ : (p : Permutation _≈_ xs ys) {iy : Any (x ≈_) ys} → - ∈-resp-↭ p (∈-resp-↭ (↭-sym p) iy) ≡ iy - ∈-resp-↭-sym⁻¹ p with eq′ ← ∈-resp-↭-sym′ (↭-sym p) - rewrite ↭-sym-involutive p = eq′ ≡.refl - --} ------------------------------------------------------------------------ -- Properties of steps, and related properties of Permutation -- previously required for proofs by well-founded induction diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/Core.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/Core.agda index bc61c2edda..764475c4ad 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/Core.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/Core.agda @@ -18,21 +18,23 @@ open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) open import Relation.Binary.Bundles using (Setoid) open import Relation.Binary.Definitions using (Reflexive; Symmetric; Transitive; LeftTrans; RightTrans) -open import Relation.Binary.Structures using (IsEquivalence; IsPreorder) +open import Relation.Binary.Structures using (IsEquivalence) -open import Data.List.Relation.Binary.Permutation.Homogeneous +open import Data.List.Relation.Binary.Permutation.Homogeneous public + using (Permutation; refl; prep; swap; trans) private variable xs ys zs : List A x y z v w : A - _≋_ _↭_ : Rel (List A) _ - _≋_ = Pointwise R - _↭_ = Permutation R - _++[_]++_ = λ xs (z : A) ys → xs ++ [ z ] ++ ys +_≋_ _↭_ : Rel (List A) _ +_≋_ = Pointwise R +_↭_ = Permutation R + + ------------------------------------------------------------------------ -- Properties of _↭_ depending on suitable assumptions on R From b4fbbf975b330c3a929c847c49e5fe8544ee892e Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:19:48 +0000 Subject: [PATCH 58/65] localised proofs of the higher-dimensional properties --- .../Properties/HigherDimensional.agda | 107 ++++++++++++++++++ 1 file changed, 107 insertions(+) create mode 100644 src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/HigherDimensional.agda diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/HigherDimensional.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/HigherDimensional.agda new file mode 100644 index 0000000000..ce4ef34b2f --- /dev/null +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties/HigherDimensional.agda @@ -0,0 +1,107 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Higher-dimensional properties of the `Homogeneous` definition of permutation +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Data.List.Relation.Binary.Permutation.Homogeneous.Properties.HigherDimensional where + +open import Data.List.Base as List using (List; []; _∷_) +open import Data.List.Relation.Binary.Pointwise as Pointwise + using (Pointwise; []; _∷_) +open import Data.List.Relation.Unary.Any as Any + using (Any; here; there) +import Data.List.Relation.Unary.Any.Properties as Any +open import Function.Base using (_∘_) +open import Level using (Level; _⊔_) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Definitions + using (Symmetric; Transitive; _Respects_) +open import Relation.Binary.PropositionalEquality.Core as ≡ + using (_≡_ ; cong; cong₂) + +import Data.List.Relation.Binary.Permutation.Homogeneous.Properties.Core as Core +import Data.List.Relation.Binary.Permutation.Homogeneous.Properties as Properties + +private + variable + a r : Level + A : Set a + x y : A + xs ys : List A + +------------------------------------------------------------------------ +-- Two higher-dimensional properties useful in the `Propositional` case, +-- specifically in the equivalence proof between `Bag` equality and `_↭_` +------------------------------------------------------------------------ + +module _ {_≈_ : Rel A r} (≈-sym : Symmetric _≈_) where + + open Core {R = _≈_} using (_≋_; _↭_; refl; prep; swap; trans; module Symmetry) + open Symmetry ≈-sym using (≋-sym; ↭-sym) + +-- involutive symmetry lifts + + module _ (≈-sym-involutive : ∀ {x y} (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) + where + + ↭-sym-involutive : (p : xs ↭ ys) → ↭-sym (↭-sym p) ≡ p + ↭-sym-involutive (refl xs≋ys) = cong refl (≋-sym-involutive xs≋ys) + where + ≋-sym-involutive : (p : xs ≋ ys) → ≋-sym (≋-sym p) ≡ p + ≋-sym-involutive [] = ≡.refl + ≋-sym-involutive (x≈y ∷ xs≋ys) rewrite ≈-sym-involutive x≈y + = cong (_ ∷_) (≋-sym-involutive xs≋ys) + ↭-sym-involutive (prep eq p) = cong₂ prep (≈-sym-involutive eq) (↭-sym-involutive p) + ↭-sym-involutive (swap eq₁ eq₂ p) rewrite ≈-sym-involutive eq₁ | ≈-sym-involutive eq₂ + = cong (swap _ _) (↭-sym-involutive p) + ↭-sym-involutive (trans p q) = cong₂ trans (↭-sym-involutive p) (↭-sym-involutive q) + + module _ (≈-trans : Transitive _≈_) where + + private + ∈-resp-Pointwise : (Any (x ≈_)) Respects _≋_ + ∈-resp-Pointwise = Properties.∈-resp-Pointwise ≈-trans + + ∈-resp-↭ : (Any (x ≈_)) Respects _↭_ + ∈-resp-↭ = Properties.∈-resp-↭ ≈-trans + +-- invertible symmetry lifts to equality on proofs of membership + + module _ (≈-trans-trans-sym : ∀ {x y z} (p : x ≈ y) (q : y ≈ z) → + ≈-trans (≈-trans p q) (≈-sym q) ≡ p) where + + ∈-resp-Pointwise-sym : (p : xs ≋ ys) {ix : Any (x ≈_) xs} → + ∈-resp-Pointwise (≋-sym p) (∈-resp-Pointwise p ix) ≡ ix + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {here ix} + = cong here (≈-trans-trans-sym ix x≈y) + ∈-resp-Pointwise-sym (x≈y ∷ xs≋ys) {there ixs} + = cong there (∈-resp-Pointwise-sym xs≋ys) + + ∈-resp-↭-sym′ : (p : ys ↭ xs) {iy : Any (x ≈_) ys} {ix : Any (x ≈_) xs} → + ix ≡ ∈-resp-↭ p iy → ∈-resp-↭ (↭-sym p) ix ≡ iy + ∈-resp-↭-sym′ (refl xs≋ys) ≡.refl = ∈-resp-Pointwise-sym xs≋ys + ∈-resp-↭-sym′ (prep eq₁ p) {here iy} {here ix} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym iy eq₁) + ∈-resp-↭-sym′ (prep eq₁ p) {there iys} {there ixs} eq + with ≡.refl ← eq = cong there (∈-resp-↭-sym′ p ≡.refl) + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {here iy} () + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {here ix} {there iys} eq + with ≡.refl ← eq = cong here (≈-trans-trans-sym ix eq₁) + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {there (here iy)} () + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (here ix)} {here iy} eq + with ≡.refl ← eq = cong (there ∘ here) (≈-trans-trans-sym ix eq₂) + ∈-resp-↭-sym′ (swap eq₁ eq₂ p) {there (there ixs)} {there (there iys)} eq + with ≡.refl ← eq = cong (there ∘ there) (∈-resp-↭-sym′ p ≡.refl) + ∈-resp-↭-sym′ (trans p₁ p₂) eq = ∈-resp-↭-sym′ p₁ (∈-resp-↭-sym′ p₂ eq) + + ∈-resp-↭-sym : (p : xs ↭ ys) {ix : Any (x ≈_) xs} → + ∈-resp-↭ (↭-sym p) (∈-resp-↭ p ix) ≡ ix + ∈-resp-↭-sym p = ∈-resp-↭-sym′ p ≡.refl + + ∈-resp-↭-sym⁻¹ : (p : xs ↭ ys) {iy : Any (x ≈_) ys} → + ∈-resp-↭ p (∈-resp-↭ (↭-sym p) iy) ≡ iy + ∈-resp-↭-sym⁻¹ p with eq′ ← ∈-resp-↭-sym′ (↭-sym p) + rewrite ↭-sym-involutive p = eq′ ≡.refl From eb0d50eef235cc326e9ca47aa534f64734435b36 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:27:36 +0000 Subject: [PATCH 59/65] drastic pruning of `import`s after refactoring --- .../Binary/Permutation/Homogeneous.agda | 38 +++++-------------- 1 file changed, 10 insertions(+), 28 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda index 5e8e379bdd..ee331d82cb 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous.agda @@ -8,37 +8,19 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where -open import Data.List.Base as List using (List; []; _∷_; [_]; _++_; length) +open import Data.List.Base as List using (List; []; _∷_) open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) -open import Data.List.Relation.Unary.All as All using (All; []; _∷_) -open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_) -open import Data.List.Relation.Unary.Any as Any using (Any; here; there) -import Data.List.Relation.Unary.Any.Properties as Any open import Data.Nat.Base using (ℕ; suc; _+_; _<_) -open import Data.Nat.Properties -open import Data.Product.Base using (_,_; _×_; ∃; ∃₂) -open import Function.Base using (_∘_; flip) open import Level using (Level; _⊔_) -open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_) -open import Relation.Binary.Bundles using (Setoid) -import Relation.Binary.Reasoning.Setoid as ≈-Reasoning -open import Relation.Binary.Structures using (IsEquivalence; IsPartialEquivalence) -open import Relation.Binary.Definitions - using ( Reflexive; Symmetric; Transitive; LeftTrans; RightTrans - ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) -open import Relation.Binary.PropositionalEquality.Core as ≡ - using (_≡_ ; cong; cong₂) -open import Relation.Nullary.Decidable using (yes; no; does) -open import Relation.Nullary.Negation using (¬_; contradiction) -open import Relation.Unary using (Pred; Decidable) +open import Relation.Binary.Core using (Rel; _⇒_) private variable - a p r s : Level - A B : Set a + a r s : Level + A : Set a xs ys zs : List A - x y z v w : A + x y x′ y′ : A ------------------------------------------------------------------------ @@ -47,11 +29,11 @@ private module _ {A : Set a} (R : Rel A r) where data Permutation : Rel (List A) (a ⊔ r) where - refl : ∀ {xs ys} → Pointwise R xs ys → Permutation xs ys - prep : ∀ {xs ys x y} (eq : R x y) → Permutation xs ys → Permutation (x ∷ xs) (y ∷ ys) - swap : ∀ {xs ys x y x′ y′} (eq₁ : R x x′) (eq₂ : R y y′) → - Permutation xs ys → Permutation (x ∷ y ∷ xs) (y′ ∷ x′ ∷ ys) - trans : Transitive Permutation + refl : Pointwise R xs ys → Permutation xs ys + prep : (eq : R x y) → Permutation xs ys → Permutation (x ∷ xs) (y ∷ ys) + swap : (eq₁ : R x x′) (eq₂ : R y y′) → Permutation xs ys → + Permutation (x ∷ y ∷ xs) (y′ ∷ x′ ∷ ys) + trans : Permutation xs ys → Permutation ys zs → Permutation xs zs ------------------------------------------------------------------------ -- Map From 014867d0fd1f9d56b4abe31619ab9ede12439e6e Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:30:18 +0000 Subject: [PATCH 60/65] rearrangement of binding prefixes after refactoring --- .../Relation/Binary/Permutation/Homogeneous/Properties.agda | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda index 4e43c50876..ce7f0ff288 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda @@ -177,7 +177,7 @@ module _ (commutativeMonoid : CommutativeMonoid a r) where foldr-commMonoid : Permutation _≈_ xs ys → foldr _∙_ ε xs ≈ foldr _∙_ ε ys foldr-commMonoid (refl xs≋ys) = Pointwise.foldr⁺ ∙-cong CM.refl xs≋ys foldr-commMonoid (prep x≈y xs↭ys) = ∙-cong x≈y (foldr-commMonoid xs↭ys) - foldr-commMonoid (swap {xs} {ys} {x} {y} {x′} {y′} x≈x′ y≈y′ xs↭ys) = begin + foldr-commMonoid (swap {x} {x′} {y} {y′} {xs} {ys} x≈x′ y≈y′ xs↭ys) = begin x ∙ (y ∙ foldr _∙_ ε xs) ≈⟨ ∙-congˡ (∙-congˡ (foldr-commMonoid xs↭ys)) ⟩ x ∙ (y ∙ foldr _∙_ ε ys) ≈⟨ assoc x y (foldr _∙_ ε ys) ⟨ (x ∙ y) ∙ foldr _∙_ ε ys ≈⟨ ∙-congʳ (comm x y) ⟩ From 9aafdf1064a39f5dfc931e05c0faa0874c3aa0c4 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:34:31 +0000 Subject: [PATCH 61/65] tightened `import`s after refactoring --- .../Relation/Binary/Permutation/Homogeneous/Properties.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda index ce7f0ff288..43cada1872 100644 --- a/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Homogeneous/Properties.agda @@ -25,7 +25,7 @@ open import Level using (Level; _⊔_) open import Relation.Binary.Core using (Rel; _Preserves_⟶_) import Relation.Binary.Reasoning.Setoid as ≈-Reasoning open import Relation.Binary.Definitions - using ( Reflexive; Symmetric; Transitive; LeftTrans; RightTrans + using ( Reflexive; Symmetric; Transitive ; _Respects_; _Respects₂_; _Respectsˡ_; _Respectsʳ_) open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_ ; cong; cong₂) @@ -245,7 +245,7 @@ module _ (sym : Symmetric S) (resp@(rʳ , rˡ) : S Respects₂ R) where ------------------------------------------------------------------------ -- Properties of steps, and related properties of Permutation -- previously required for proofs by well-founded induction --- rendered obsolete/deprecatable by ↭-transˡ-≋ , ↭-transʳ-≋ +-- rendered obsolete/deprecatable by Core.↭-transˡ-≋ , Core.↭-transʳ-≋ ------------------------------------------------------------------------ module Steps {R : Rel A r} where From dd451bc9a85ef7530c40c5083aa3b07fb323ca79 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 14:50:32 +0000 Subject: [PATCH 62/65] follow through --- .../Binary/Permutation/Propositional.agda | 3 ++- .../Relation/Binary/Permutation/Setoid.agda | 20 ++++++++++--------- 2 files changed, 13 insertions(+), 10 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Propositional.agda b/src/Data/List/Relation/Binary/Permutation/Propositional.agda index 5efbbe76a1..a2e12f1409 100644 --- a/src/Data/List/Relation/Binary/Permutation/Propositional.agda +++ b/src/Data/List/Relation/Binary/Permutation/Propositional.agda @@ -19,6 +19,7 @@ open import Relation.Binary.Definitions open import Relation.Binary.PropositionalEquality using (_≡_; refl; setoid) import Data.List.Relation.Binary.Permutation.Setoid as Permutation +import Data.List.Relation.Binary.Permutation.Homogeneous.Properties.Core as Core ------------------------------------------------------------------------ -- Definition of permutation @@ -50,7 +51,7 @@ open ↭ public ↭-transʳ-≋ xs↭ys ys≋zs with refl ← ≋⇒≡ ys≋zs = xs↭ys ↭-trans : Transitive _↭_ -↭-trans = ↭-trans′ ↭-transˡ-≋ ↭-transʳ-≋ +↭-trans = Core.LRTransitivity.↭-trans ↭-transˡ-≋ ↭-transʳ-≋ ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid.agda b/src/Data/List/Relation/Binary/Permutation/Setoid.agda index 9cc5fdab10..f4513a9bf8 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid.agda @@ -5,6 +5,7 @@ ------------------------------------------------------------------------ {-# OPTIONS --cubical-compatible --safe #-} +{-# OPTIONS --warn=noUserWarning #-} -- for deprecated function `steps` open import Function.Base using (_∘′_) open import Relation.Binary.Core using (Rel; _⇒_) @@ -19,6 +20,7 @@ module Data.List.Relation.Binary.Permutation.Setoid open import Data.List.Base using (List; _∷_) import Data.List.Relation.Binary.Permutation.Homogeneous as Homogeneous +import Data.List.Relation.Binary.Permutation.Homogeneous.Properties.Core as Core open import Data.List.Relation.Binary.Equality.Setoid S open import Data.Nat.Base using (ℕ) open import Level using (_⊔_) @@ -35,7 +37,7 @@ open Setoid S infix 3 _↭_ _↭_ : Rel (List A) (a ⊔ ℓ) -_↭_ = Homogeneous.Permutation _≈_ +_↭_ = Core._↭_ {R = _≈_} ------------------------------------------------------------------------ -- Smart constructor aliases @@ -45,37 +47,37 @@ _↭_ = Homogeneous.Permutation _≈_ -- prepended are *propositionally* equal ↭-pointwise : _≋_ ⇒ _↭_ -↭-pointwise = Homogeneous.↭-pointwise +↭-pointwise = Core.↭-pointwise ↭-prep : ∀ x {xs ys} → xs ↭ ys → x ∷ xs ↭ x ∷ ys -↭-prep _ = Homogeneous.↭-prep ≈-refl +↭-prep _ = Core.Reflexivity.↭-prep ≈-refl ↭-swap : ∀ x y {xs ys} → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys -↭-swap = Homogeneous.↭-swap ≈-refl +↭-swap = Core.Reflexivity.↭-swap ≈-refl ↭-trans′ : LeftTrans _≋_ _↭_ → RightTrans _↭_ _≋_ → Transitive _↭_ -↭-trans′ = Homogeneous.↭-trans′ +↭-trans′ = Core.LRTransitivity.↭-trans ------------------------------------------------------------------------ -- Functions over permutations (retained for legacy) steps : ∀ {xs ys} → xs ↭ ys → ℕ -steps = Homogeneous.Steps.steps +steps = Homogeneous.steps ------------------------------------------------------------------------ -- _↭_ is an equivalence ↭-reflexive : _≡_ ⇒ _↭_ -↭-reflexive refl = Homogeneous.↭-refl′ ≈-refl +↭-reflexive refl = Core.Reflexivity.↭-refl ≈-refl ↭-refl : Reflexive _↭_ ↭-refl = ↭-reflexive refl ↭-sym : Symmetric _↭_ -↭-sym = Homogeneous.↭-sym′ ≈-sym +↭-sym = Core.Symmetry.↭-sym ≈-sym ↭-trans : Transitive _↭_ -↭-trans = Homogeneous.↭-trans ≈-trans +↭-trans = Core.Transitivity.↭-trans ≈-trans ↭-isEquivalence : IsEquivalence _↭_ ↭-isEquivalence = record { refl = ↭-refl ; sym = ↭-sym ; trans = ↭-trans } From 55276e8d59627709286e4561705f076d5bb78152 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Tue, 26 Mar 2024 15:07:14 +0000 Subject: [PATCH 63/65] fix up parametrisation (first go) --- .../Binary/Permutation/Setoid/Properties.agda | 53 ++++++++++--------- 1 file changed, 29 insertions(+), 24 deletions(-) diff --git a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda index ad7c50279f..7e3f9b3341 100644 --- a/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda +++ b/src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda @@ -40,7 +40,8 @@ open import Relation.Unary using (Pred; Decidable) import Data.List.Relation.Binary.Equality.Setoid as Equality import Data.List.Relation.Binary.Permutation.Setoid as Permutation -import Data.List.Relation.Binary.Permutation.Homogeneous as Properties +import Data.List.Relation.Binary.Permutation.Homogeneous.Properties as Properties +--import Data.List.Relation.Binary.Permutation.Homogeneous.Properties.Core as Core open Setoid S using (_≈_; isEquivalence) @@ -92,11 +93,11 @@ Unique-resp-↭ = AllPairs-resp-↭ (_∘ ≈-sym) ≉-resp₂ ↭-respˡ-≋ = Properties.Steps.↭-respˡ-≋ ≈-trans ≈-sym ↭-transˡ-≋ : LeftTrans _≋_ _↭_ -↭-transˡ-≋ = Properties.↭-transˡ-≋ ≈-trans +↭-transˡ-≋ = Properties.Transitivity.↭-transˡ-≋ ≈-trans ↭-transʳ-≋ : RightTrans _↭_ _≋_ -↭-transʳ-≋ = Properties.↭-transʳ-≋ ≈-trans - +↭-transʳ-≋ = Properties.Transitivity.↭-transʳ-≋ ≈-trans +{- module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym p) ≡ p) where @@ -116,11 +117,11 @@ module _ (≈-sym-involutive : ∀ {x y} → (p : x ≈ y) → ≈-sym (≈-sym ∈-resp-↭ (↭-sym p) (∈-resp-↭ p ix) ≡ ix ∈-resp-↭-sym = Properties.∈-resp-↭-sym ≈-trans ≈-sym ≈-sym-involutive ≈-trans-trans-sym - +-} ------------------------------------------------------------------------ -- Properties of steps (legacy) ------------------------------------------------------------------------ - +{- 0 Date: Wed, 27 Mar 2024 09:45:53 +0000 Subject: [PATCH 64/65] =?UTF-8?q?reverted=20use=20of=20`=E2=86=AD-shift`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../List/Relation/Ternary/Interleaving/Propositional.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda index 60cec992fd..a93abb7ee6 100644 --- a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda +++ b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda @@ -10,7 +10,7 @@ module Data.List.Relation.Ternary.Interleaving.Propositional {a} {A : Set a} whe open import Data.List.Base as List using (List; []; _∷_; _++_) open import Data.List.Relation.Binary.Permutation.Propositional as Perm using (_↭_) -open import Data.List.Relation.Binary.Permutation.Propositional.Properties using (↭-shift) +open import Data.List.Relation.Binary.Permutation.Propositional.Properties using (shift) import Data.List.Relation.Ternary.Interleaving.Setoid as General open import Relation.Binary.PropositionalEquality.Core using (refl) open import Relation.Binary.PropositionalEquality.Properties using (setoid) @@ -38,5 +38,5 @@ toPermutation [] = Perm.↭-refl toPermutation (consˡ sp) = Perm.↭-prep _ (toPermutation sp) toPermutation {l} {r ∷ rs} {a ∷ as} (consʳ sp) = begin a ∷ as <⟨ toPermutation sp ⟩ - a ∷ l ++ rs ↭⟨ ↭-shift l ⟨ + a ∷ l ++ rs ↭⟨ shift a l rs ⟨ l ++ a ∷ rs ∎ From a07557cbcbb2aed60fcc58f0a58c64e0e695c5d9 Mon Sep 17 00:00:00 2001 From: James McKinna Date: Wed, 27 Mar 2024 09:47:40 +0000 Subject: [PATCH 65/65] =?UTF-8?q?reverted=20use=20of=20`=E2=86=AD--refl`?= =?UTF-8?q?=20and=20=20`=E2=86=AD--prep`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../List/Relation/Ternary/Interleaving/Propositional.agda | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda index a93abb7ee6..fd198e20f2 100644 --- a/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda +++ b/src/Data/List/Relation/Ternary/Interleaving/Propositional.agda @@ -34,8 +34,8 @@ pattern consʳ xs = refl ∷ʳ xs -- New combinators toPermutation : ∀ {l r as} → Interleaving l r as → as ↭ l ++ r -toPermutation [] = Perm.↭-refl -toPermutation (consˡ sp) = Perm.↭-prep _ (toPermutation sp) +toPermutation [] = Perm.refl +toPermutation (consˡ sp) = Perm.prep _ (toPermutation sp) toPermutation {l} {r ∷ rs} {a ∷ as} (consʳ sp) = begin a ∷ as <⟨ toPermutation sp ⟩ a ∷ l ++ rs ↭⟨ shift a l rs ⟨