Skip to content

Commit

Permalink
Refactor Data.List.Relation.Binary.Permutation.*, part I (#2333)
Browse files Browse the repository at this point in the history
* move `steps` towards deprecation in `Homogeneous`

* deprecate `steps`; refactor `Setoid` proofs and equaiotnal reasoning combinators

* extensive refactoring

* tidy up

* add equivalence with `Setoid` representation

* removed buggy `PermutationReasoning` syntax

* refactored; removed buggy `PermutationReasoning` syntax

* `CHANGELOG`

* final fix-ups

* tighten `import`s

* tighten `import`s

* redundant constructor aliases

* fix-up `Reasoning` steps with the alias

* use aliases

* more `import` tightening

* refactor: encapsulate and tighten up

* avoid `PermutationReasoning` custom combinators

* fix up `CHANGELOG`

* encapsulate helper function

* revert changes

* review comments

* `fix-whitespace`

* toned down the comment on `steps`

* remove use of infix `insert`

* revert other deprecation

* no need for qualification

* remove deprecation banner

* three paras of commentary on the new transitivity proofs

* missing entry

* missing terminator

* response to review comments

* `with` to `let`

* fixed `BagAndSetEquality`

* fixed qualified `import` bug introduced during merge conflict resolution

* Update CHANGELOG.md

Deleted spurious attribution of the lemmas in `Data.List.Properties` about `product` to `Data.List.Relation.Unary.All.Properties`. Hope this fixes things now!

---------

Co-authored-by: MatthewDaggitt <[email protected]>
  • Loading branch information
jamesmckinna and MatthewDaggitt authored Sep 28, 2024
1 parent d13d1ef commit 45a46f7
Show file tree
Hide file tree
Showing 7 changed files with 551 additions and 431 deletions.
60 changes: 55 additions & 5 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -59,6 +59,13 @@ Deprecated names
normalise-correct ↦ Algebra.Solver.Monoid.Normal.correct
```

* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`:
```agda
split ↦ ↭-split
```
with a more informative type (see below).
```
* In `Data.Vec.Properties`:
```agda
++-assoc _ ↦ ++-assoc-eqFree
Expand Down Expand Up @@ -167,11 +174,6 @@ Additions to existing modules
concatMap-++ : concatMap f (xs ++ ys) ≡ concatMap f xs ++ concatMap f ys
```

* In `Data.List.Relation.Unary.All`:
```agda
search : Decidable P → ∀ xs → All (∁ P) xs ⊎ Any P xs
```

* In `Data.List.Relation.Unary.Any.Properties`:
```agda
concatMap⁺ : Any (Any P ∘ f) xs → Any P (concatMap f xs)
Expand All @@ -194,12 +196,60 @@ Additions to existing modules
++⁺ˡ : ∀ zs → ws ≋ xs → ws ++ zs ≋ xs ++ zs
```

* In `Data.List.Relation.Binary.Permutation.Homogeneous`:
```agda
steps : Permutation R xs ys → ℕ
```

* In `Data.List.Relation.Binary.Permutation.Propositional`:
constructor aliases
```agda
↭-refl : Reflexive _↭_
↭-prep : ∀ x → xs ↭ ys → x ∷ xs ↭ x ∷ ys
↭-swap : ∀ x y → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys
```
and properties
```agda
↭-reflexive-≋ : _≋_ ⇒ _↭_
↭⇒↭ₛ : _↭_ ⇒ _↭ₛ_
↭ₛ⇒↭ : _↭ₛ_ ⇒ _↭_
```
where `_↭ₛ_` is the `Setoid (setoid _)` instance of `Permutation`

* In `Data.List.Relation.Binary.Permutation.Propositional.Properties`:
```agda
Any-resp-[σ∘σ⁻¹] : (σ : xs ↭ ys) (iy : Any P ys) →
Any-resp-↭ (trans (↭-sym σ) σ) iy ≡ iy
∈-resp-[σ∘σ⁻¹] : (σ : xs ↭ ys) (iy : y ∈ ys) →
∈-resp-↭ (trans (↭-sym σ) σ) iy ≡ iy
product-↭ : product Preserves _↭_ ⟶ _≡_
```

* In `Data.List.Relation.Binary.Permutation.Setoid`:
```agda
↭-reflexive-≋ : _≋_ ⇒ _↭_
↭-transˡ-≋ : LeftTrans _≋_ _↭_
↭-transʳ-≋ : RightTrans _↭_ _≋_
↭-trans′ : Transitive _↭_
```

* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`:
```agda
↭-split : xs ↭ (as ++ [ v ] ++ bs) →
∃₂ λ ps qs → xs ≋ (ps ++ [ v ] ++ qs) × (ps ++ qs) ↭ (as ++ bs)
drop-∷ : x ∷ xs ↭ x ∷ ys → xs ↭ ys
```

* In `Data.List.Relation.Binary.Pointwise`:
```agda
++⁺ʳ : Reflexive R → ∀ xs → (xs ++_) Preserves (Pointwise R) ⟶ (Pointwise R)
++⁺ˡ : Reflexive R → ∀ zs → (_++ zs) Preserves (Pointwise R) ⟶ (Pointwise R)
```

* In `Data.List.Relation.Unary.All`:
```agda
search : Decidable P → ∀ xs → All (∁ P) xs ⊎ Any P xs
* In `Data.List.Relation.Binary.Subset.Setoid.Properties`:
```agda
∷⊈[] : x ∷ xs ⊈ []
Expand Down
29 changes: 9 additions & 20 deletions src/Data/List/Relation/Binary/BagAndSetEquality.agda
Original file line number Diff line number Diff line change
Expand Up @@ -27,9 +27,9 @@ 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
using (_↭_; ↭-sym; refl; module PermutationReasoning)
using (_↭_; ↭-refl; ↭-sym; ↭-prep; module PermutationReasoning)
open import Data.List.Relation.Binary.Permutation.Propositional.Properties
using (∈-resp-↭; ∈-resp-[σ⁻¹∘σ]; ↭-sym-involutive; shift; ++-comm)
using (∈-resp-↭; ∈-resp-[σ⁻¹∘σ]; ∈-resp-[σ∘σ⁻¹]; shift; ++-comm)
open import Data.Product.Base as Product using (∃; _,_; proj₁; proj₂; _×_)
import Data.Product.Function.Dependent.Propositional as Σ
open import Data.Sum.Base as Sum using (_⊎_; [_,_]′; inj₁; inj₂)
Expand Down Expand Up @@ -574,29 +574,18 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys =


------------------------------------------------------------------------
-- Relationships to other relations
-- Relationships to propositional permutation

↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_
↭⇒∼bag 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 = ∈-resp-↭ xs↭ys

from : {xs ys} xs ↭ ys v ∈ ys v ∈ xs
from xs↭ys = ∈-resp-↭ (↭-sym xs↭ys)

from∘to : {xs ys} (p : xs ↭ ys) (q : v ∈ xs) from p (to p q) ≡ q
from∘to = ∈-resp-[σ⁻¹∘σ]

to∘from : {xs ys} (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 xs↭ys {v} =
mk↔ₛ′ (∈-resp-↭ xs↭ys) (∈-resp-↭ (↭-sym xs↭ys)) (∈-resp-[σ∘σ⁻¹] xs↭ys) (∈-resp-[σ⁻¹∘σ] xs↭ys)

∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A}
∼bag⇒↭ {A = A} {[]} eq with refl empty-unique (↔-sym eq) = refl
∼bag⇒↭ {A = A} {[]} eq with refl empty-unique (↔-sym eq) = ↭-refl
∼bag⇒↭ {A = A} {x ∷ xs} eq
with zs₁ , zs₂ , p ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) rewrite p = begin
x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩
x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩
with zs₁ , zs₂ , refl ∈-∃++ (Inverse.to (eq {x}) (here refl)) = begin
x ∷ xs ↭⟨ ↭-prep x (∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂))))) ⟩
x ∷ (zs₂ ++ zs₁) ↭⟨ ↭-prep x (++-comm zs₂ zs₁)
x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨
zs₁ ++ x ∷ zs₂ ∎
where
Expand Down
12 changes: 10 additions & 2 deletions src/Data/List/Relation/Binary/Permutation/Homogeneous.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,8 +11,8 @@ 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)
import Data.List.Relation.Binary.Pointwise.Properties as Pointwise
open import Data.Nat.Base using (ℕ; suc; _+_)
open import Level using (Level; _⊔_)
open import Relation.Binary.Core using (Rel; _⇒_)
open import Relation.Binary.Bundles using (Setoid)
Expand Down Expand Up @@ -59,3 +59,11 @@ 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)

-- Measures the number of constructors, can be useful for termination proofs

steps : {R : Rel A r} {xs ys} 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
65 changes: 53 additions & 12 deletions src/Data/List/Relation/Binary/Permutation/Propositional.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,14 +10,22 @@ 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)
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl)
import Relation.Binary.Reasoning.Setoid as EqReasoning
open import Relation.Binary.Reasoning.Syntax

import Data.List.Relation.Binary.Permutation.Setoid as Permutation

private
variable
x y z v w : A
ws xs ys zs : List A

------------------------------------------------------------------------
-- An inductive definition of permutation

Expand All @@ -30,21 +38,32 @@ open import Relation.Binary.Reasoning.Syntax
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
refl : xs ↭ xs
prep : x xs ↭ ys x ∷ xs ↭ x ∷ ys
swap : x y xs ↭ ys x ∷ y ∷ xs ↭ y ∷ x ∷ ys
trans : xs ↭ ys ys ↭ zs xs ↭ zs

-- Constructor aliases

↭-refl : Reflexive _↭_
↭-refl = refl

↭-prep : x xs ↭ ys x ∷ xs ↭ x ∷ ys
↭-prep = prep

↭-swap : x y xs ↭ ys x ∷ y ∷ xs ↭ y ∷ x ∷ ys
↭-swap = swap

------------------------------------------------------------------------
-- _↭_ is an equivalence

↭-reflexive : _≡_ ⇒ _↭_
↭-reflexive refl = refl
↭-reflexive refl = ↭-refl

↭-refl : Reflexive _↭_
↭-refl = refl
↭-reflexive-≋ : _≋_ ⇒ _↭_
↭-reflexive-≋ xs≋ys = ↭-reflexive (≋⇒≡ xs≋ys)

↭-sym : {xs ys} xs ↭ ys ys ↭ xs
↭-sym : 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)
Expand All @@ -58,7 +77,7 @@ data _↭_ : Rel (List A) a where

↭-isEquivalence : IsEquivalence _↭_
↭-isEquivalence = record
{ refl = refl
{ refl = ↭-refl
; sym = ↭-sym
; trans = ↭-trans
}
Expand All @@ -68,6 +87,28 @@ data _↭_ : Rel (List A) a where
{ isEquivalence = ↭-isEquivalence
}

------------------------------------------------------------------------
-- _↭_ is equivalent to `Setoid`-based permutation

private
open module ↭ₛ = Permutation (≡.setoid A)
using ()
renaming (_↭_ to _↭ₛ_)

↭⇒↭ₛ : xs ↭ ys xs ↭ₛ ys
↭⇒↭ₛ refl = ↭ₛ.↭-refl
↭⇒↭ₛ (prep x p) = ↭ₛ.↭-prep x (↭⇒↭ₛ p)
↭⇒↭ₛ (swap x y p) = ↭ₛ.↭-swap x y (↭⇒↭ₛ p)
↭⇒↭ₛ (trans p q) = ↭ₛ.↭-trans′ (↭⇒↭ₛ p) (↭⇒↭ₛ q)


↭ₛ⇒↭ : _↭ₛ_ ⇒ _↭_
↭ₛ⇒↭ (↭ₛ.refl xs≋ys) = ↭-reflexive-≋ xs≋ys
↭ₛ⇒↭ (↭ₛ.prep refl p) = ↭-prep _ (↭ₛ⇒↭ p)
↭ₛ⇒↭ (↭ₛ.swap refl refl p) = ↭-swap _ _ (↭ₛ⇒↭ p)
↭ₛ⇒↭ (↭ₛ.trans p q) = ↭-trans (↭ₛ⇒↭ p) (↭ₛ⇒↭ q)


------------------------------------------------------------------------
-- A reasoning API to chain permutation proofs and allow "zooming in"
-- to localised reasoning.
Expand All @@ -89,12 +130,12 @@ module PermutationReasoning where
-- 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))
step-prep x xs rel xs↭ys = ↭-go (↭-prep x xs↭ys) 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))
step-swap x y xs rel xs↭ys = ↭-go (↭-swap x y xs↭ys) 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
Loading

0 comments on commit 45a46f7

Please sign in to comment.