Skip to content

Commit 45a46f7

Browse files
Refactor Data.List.Relation.Binary.Permutation.*, part I (#2333)
* 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]>
1 parent d13d1ef commit 45a46f7

File tree

7 files changed

+551
-431
lines changed

7 files changed

+551
-431
lines changed

CHANGELOG.md

Lines changed: 55 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -59,6 +59,13 @@ Deprecated names
5959
normalise-correct ↦ Algebra.Solver.Monoid.Normal.correct
6060
```
6161

62+
* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`:
63+
```agda
64+
split ↦ ↭-split
65+
```
66+
with a more informative type (see below).
67+
```
68+
6269
* In `Data.Vec.Properties`:
6370
```agda
6471
++-assoc _ ↦ ++-assoc-eqFree
@@ -167,11 +174,6 @@ Additions to existing modules
167174
concatMap-++ : concatMap f (xs ++ ys) ≡ concatMap f xs ++ concatMap f ys
168175
```
169176

170-
* In `Data.List.Relation.Unary.All`:
171-
```agda
172-
search : Decidable P → ∀ xs → All (∁ P) xs ⊎ Any P xs
173-
```
174-
175177
* In `Data.List.Relation.Unary.Any.Properties`:
176178
```agda
177179
concatMap⁺ : Any (Any P ∘ f) xs → Any P (concatMap f xs)
@@ -194,12 +196,60 @@ Additions to existing modules
194196
++⁺ˡ : ∀ zs → ws ≋ xs → ws ++ zs ≋ xs ++ zs
195197
```
196198

199+
* In `Data.List.Relation.Binary.Permutation.Homogeneous`:
200+
```agda
201+
steps : Permutation R xs ys → ℕ
202+
```
203+
204+
* In `Data.List.Relation.Binary.Permutation.Propositional`:
205+
constructor aliases
206+
```agda
207+
↭-refl : Reflexive _↭_
208+
↭-prep : ∀ x → xs ↭ ys → x ∷ xs ↭ x ∷ ys
209+
↭-swap : ∀ x y → xs ↭ ys → x ∷ y ∷ xs ↭ y ∷ x ∷ ys
210+
```
211+
and properties
212+
```agda
213+
↭-reflexive-≋ : _≋_ ⇒ _↭_
214+
↭⇒↭ₛ : _↭_ ⇒ _↭ₛ_
215+
↭ₛ⇒↭ : _↭ₛ_ ⇒ _↭_
216+
```
217+
where `_↭ₛ_` is the `Setoid (setoid _)` instance of `Permutation`
218+
219+
* In `Data.List.Relation.Binary.Permutation.Propositional.Properties`:
220+
```agda
221+
Any-resp-[σ∘σ⁻¹] : (σ : xs ↭ ys) (iy : Any P ys) →
222+
Any-resp-↭ (trans (↭-sym σ) σ) iy ≡ iy
223+
∈-resp-[σ∘σ⁻¹] : (σ : xs ↭ ys) (iy : y ∈ ys) →
224+
∈-resp-↭ (trans (↭-sym σ) σ) iy ≡ iy
225+
product-↭ : product Preserves _↭_ ⟶ _≡_
226+
```
227+
228+
* In `Data.List.Relation.Binary.Permutation.Setoid`:
229+
```agda
230+
↭-reflexive-≋ : _≋_ ⇒ _↭_
231+
↭-transˡ-≋ : LeftTrans _≋_ _↭_
232+
↭-transʳ-≋ : RightTrans _↭_ _≋_
233+
↭-trans′ : Transitive _↭_
234+
```
235+
236+
* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`:
237+
```agda
238+
↭-split : xs ↭ (as ++ [ v ] ++ bs) →
239+
∃₂ λ ps qs → xs ≋ (ps ++ [ v ] ++ qs) × (ps ++ qs) ↭ (as ++ bs)
240+
drop-∷ : x ∷ xs ↭ x ∷ ys → xs ↭ ys
241+
```
242+
197243
* In `Data.List.Relation.Binary.Pointwise`:
198244
```agda
199245
++⁺ʳ : Reflexive R → ∀ xs → (xs ++_) Preserves (Pointwise R) ⟶ (Pointwise R)
200246
++⁺ˡ : Reflexive R → ∀ zs → (_++ zs) Preserves (Pointwise R) ⟶ (Pointwise R)
201247
```
202248

249+
* In `Data.List.Relation.Unary.All`:
250+
```agda
251+
search : Decidable P → ∀ xs → All (∁ P) xs ⊎ Any P xs
252+
203253
* In `Data.List.Relation.Binary.Subset.Setoid.Properties`:
204254
```agda
205255
∷⊈[] : x ∷ xs ⊈ []

src/Data/List/Relation/Binary/BagAndSetEquality.agda

Lines changed: 9 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -27,9 +27,9 @@ open import Data.List.Membership.Propositional.Properties
2727
open import Data.List.Relation.Binary.Subset.Propositional.Properties
2828
using (⊆-preorder)
2929
open import Data.List.Relation.Binary.Permutation.Propositional
30-
using (_↭_; ↭-sym; refl; module PermutationReasoning)
30+
using (_↭_; ↭-refl; ↭-sym; ↭-prep; module PermutationReasoning)
3131
open import Data.List.Relation.Binary.Permutation.Propositional.Properties
32-
using (∈-resp-↭; ∈-resp-[σ⁻¹∘σ]; ↭-sym-involutive; shift; ++-comm)
32+
using (∈-resp-↭; ∈-resp-[σ⁻¹∘σ]; ∈-resp-[σ∘σ⁻¹]; shift; ++-comm)
3333
open import Data.Product.Base as Product using (∃; _,_; proj₁; proj₂; _×_)
3434
import Data.Product.Function.Dependent.Propositional as Σ
3535
open import Data.Sum.Base as Sum using (_⊎_; [_,_]′; inj₁; inj₂)
@@ -574,29 +574,18 @@ drop-cons {x = x} {xs} {ys} x∷xs≈x∷ys =
574574

575575

576576
------------------------------------------------------------------------
577-
-- Relationships to other relations
577+
-- Relationships to propositional permutation
578578

579579
↭⇒∼bag : _↭_ {A = A} ⇒ _∼[ bag ]_
580-
↭⇒∼bag xs↭ys {v} = mk↔ₛ′ (to xs↭ys) (from xs↭ys) (to∘from xs↭ys) (from∘to xs↭ys)
581-
where
582-
to : {xs ys} xs ↭ ys v ∈ xs v ∈ ys
583-
to xs↭ys = ∈-resp-↭ xs↭ys
584-
585-
from : {xs ys} xs ↭ ys v ∈ ys v ∈ xs
586-
from xs↭ys = ∈-resp-↭ (↭-sym xs↭ys)
587-
588-
from∘to : {xs ys} (p : xs ↭ ys) (q : v ∈ xs) from p (to p q) ≡ q
589-
from∘to = ∈-resp-[σ⁻¹∘σ]
590-
591-
to∘from : {xs ys} (p : xs ↭ ys) (q : v ∈ ys) to p (from p q) ≡ q
592-
to∘from p with res from∘to (↭-sym p) rewrite ↭-sym-involutive p = res
580+
↭⇒∼bag xs↭ys {v} =
581+
mk↔ₛ′ (∈-resp-↭ xs↭ys) (∈-resp-↭ (↭-sym xs↭ys)) (∈-resp-[σ∘σ⁻¹] xs↭ys) (∈-resp-[σ⁻¹∘σ] xs↭ys)
593582

594583
∼bag⇒↭ : _∼[ bag ]_ ⇒ _↭_ {A = A}
595-
∼bag⇒↭ {A = A} {[]} eq with refl empty-unique (↔-sym eq) = refl
584+
∼bag⇒↭ {A = A} {[]} eq with refl empty-unique (↔-sym eq) = ↭-refl
596585
∼bag⇒↭ {A = A} {x ∷ xs} eq
597-
with zs₁ , zs₂ , p ∈-∃++ (Inverse.to (eq {x}) (here ≡.refl)) rewrite p = begin
598-
x ∷ xs <⟨ ∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂)))) ⟩
599-
x ∷ (zs₂ ++ zs₁) <⟨ ++-comm zs₂ zs₁ ⟩
586+
with zs₁ , zs₂ , refl ∈-∃++ (Inverse.to (eq {x}) (here refl)) = begin
587+
x ∷ xs ↭⟨ ↭-prep x (∼bag⇒↭ (drop-cons (↔-trans eq (comm zs₁ (x ∷ zs₂))))) ⟩
588+
x ∷ (zs₂ ++ zs₁) ↭⟨ ↭-prep x (++-comm zs₂ zs₁)
600589
x ∷ (zs₁ ++ zs₂) ↭⟨ shift x zs₁ zs₂ ⟨
601590
zs₁ ++ x ∷ zs₂ ∎
602591
where

src/Data/List/Relation/Binary/Permutation/Homogeneous.agda

Lines changed: 10 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -11,8 +11,8 @@ module Data.List.Relation.Binary.Permutation.Homogeneous where
1111
open import Data.List.Base using (List; _∷_)
1212
open import Data.List.Relation.Binary.Pointwise.Base as Pointwise
1313
using (Pointwise)
14-
open import Data.List.Relation.Binary.Pointwise.Properties as Pointwise
15-
using (symmetric)
14+
import Data.List.Relation.Binary.Pointwise.Properties as Pointwise
15+
open import Data.Nat.Base using (ℕ; suc; _+_)
1616
open import Level using (Level; _⊔_)
1717
open import Relation.Binary.Core using (Rel; _⇒_)
1818
open import Relation.Binary.Bundles using (Setoid)
@@ -59,3 +59,11 @@ map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys)
5959
map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys)
6060
map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys)
6161
map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs)
62+
63+
-- Measures the number of constructors, can be useful for termination proofs
64+
65+
steps : {R : Rel A r} {xs ys} Permutation R xs ys
66+
steps (refl _) = 1
67+
steps (prep _ xs↭ys) = suc (steps xs↭ys)
68+
steps (swap _ _ xs↭ys) = suc (steps xs↭ys)
69+
steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs

src/Data/List/Relation/Binary/Permutation/Propositional.agda

Lines changed: 53 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -10,14 +10,22 @@ module Data.List.Relation.Binary.Permutation.Propositional
1010
{a} {A : Set a} where
1111

1212
open import Data.List.Base using (List; []; _∷_)
13+
open import Data.List.Relation.Binary.Equality.Propositional using (_≋_; ≋⇒≡)
1314
open import Relation.Binary.Core using (Rel; _⇒_)
1415
open import Relation.Binary.Bundles using (Setoid)
1516
open import Relation.Binary.Structures using (IsEquivalence)
1617
open import Relation.Binary.Definitions using (Reflexive; Transitive)
17-
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
18+
open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl)
1819
import Relation.Binary.Reasoning.Setoid as EqReasoning
1920
open import Relation.Binary.Reasoning.Syntax
2021

22+
import Data.List.Relation.Binary.Permutation.Setoid as Permutation
23+
24+
private
25+
variable
26+
x y z v w : A
27+
ws xs ys zs : List A
28+
2129
------------------------------------------------------------------------
2230
-- An inductive definition of permutation
2331

@@ -30,21 +38,32 @@ open import Relation.Binary.Reasoning.Syntax
3038
infix 3 _↭_
3139

3240
data _↭_ : Rel (List A) a where
33-
refl : {xs} xs ↭ xs
34-
prep : {xs ys} x xs ↭ ys x ∷ xs ↭ x ∷ ys
35-
swap : {xs ys} x y xs ↭ ys x ∷ y ∷ xs ↭ y ∷ x ∷ ys
36-
trans : {xs ys zs} xs ↭ ys ys ↭ zs xs ↭ zs
41+
refl : xs ↭ xs
42+
prep : x xs ↭ ys x ∷ xs ↭ x ∷ ys
43+
swap : x y xs ↭ ys x ∷ y ∷ xs ↭ y ∷ x ∷ ys
44+
trans : xs ↭ ys ys ↭ zs xs ↭ zs
45+
46+
-- Constructor aliases
47+
48+
↭-refl : Reflexive _↭_
49+
↭-refl = refl
50+
51+
↭-prep : x xs ↭ ys x ∷ xs ↭ x ∷ ys
52+
↭-prep = prep
53+
54+
↭-swap : x y xs ↭ ys x ∷ y ∷ xs ↭ y ∷ x ∷ ys
55+
↭-swap = swap
3756

3857
------------------------------------------------------------------------
3958
-- _↭_ is an equivalence
4059

4160
↭-reflexive : _≡_ ⇒ _↭_
42-
↭-reflexive refl = refl
61+
↭-reflexive refl = ↭-refl
4362

44-
↭-refl : Reflexive _↭_
45-
↭-refl = refl
63+
↭-reflexive-≋ : _≋_ ⇒ _↭_
64+
↭-reflexive-≋ xs≋ys = ↭-reflexive (≋⇒≡ xs≋ys)
4665

47-
↭-sym : {xs ys} xs ↭ ys ys ↭ xs
66+
↭-sym : xs ↭ ys ys ↭ xs
4867
↭-sym refl = refl
4968
↭-sym (prep x xs↭ys) = prep x (↭-sym xs↭ys)
5069
↭-sym (swap x y xs↭ys) = swap y x (↭-sym xs↭ys)
@@ -58,7 +77,7 @@ data _↭_ : Rel (List A) a where
5877

5978
↭-isEquivalence : IsEquivalence _↭_
6079
↭-isEquivalence = record
61-
{ refl = refl
80+
{ refl = ↭-refl
6281
; sym = ↭-sym
6382
; trans = ↭-trans
6483
}
@@ -68,6 +87,28 @@ data _↭_ : Rel (List A) a where
6887
{ isEquivalence = ↭-isEquivalence
6988
}
7089

90+
------------------------------------------------------------------------
91+
-- _↭_ is equivalent to `Setoid`-based permutation
92+
93+
private
94+
open module ↭ₛ = Permutation (≡.setoid A)
95+
using ()
96+
renaming (_↭_ to _↭ₛ_)
97+
98+
↭⇒↭ₛ : xs ↭ ys xs ↭ₛ ys
99+
↭⇒↭ₛ refl = ↭ₛ.↭-refl
100+
↭⇒↭ₛ (prep x p) = ↭ₛ.↭-prep x (↭⇒↭ₛ p)
101+
↭⇒↭ₛ (swap x y p) = ↭ₛ.↭-swap x y (↭⇒↭ₛ p)
102+
↭⇒↭ₛ (trans p q) = ↭ₛ.↭-trans′ (↭⇒↭ₛ p) (↭⇒↭ₛ q)
103+
104+
105+
↭ₛ⇒↭ : _↭ₛ_ ⇒ _↭_
106+
↭ₛ⇒↭ (↭ₛ.refl xs≋ys) = ↭-reflexive-≋ xs≋ys
107+
↭ₛ⇒↭ (↭ₛ.prep refl p) = ↭-prep _ (↭ₛ⇒↭ p)
108+
↭ₛ⇒↭ (↭ₛ.swap refl refl p) = ↭-swap _ _ (↭ₛ⇒↭ p)
109+
↭ₛ⇒↭ (↭ₛ.trans p q) = ↭-trans (↭ₛ⇒↭ p) (↭ₛ⇒↭ q)
110+
111+
71112
------------------------------------------------------------------------
72113
-- A reasoning API to chain permutation proofs and allow "zooming in"
73114
-- to localised reasoning.
@@ -89,12 +130,12 @@ module PermutationReasoning where
89130
-- Skip reasoning on the first element
90131
step-prep : x xs {ys zs : List A} (x ∷ ys) IsRelatedTo zs
91132
xs ↭ ys (x ∷ xs) IsRelatedTo zs
92-
step-prep x xs rel xs↭ys = relTo (trans (prep x xs↭ys) (begin rel))
133+
step-prep x xs rel xs↭ys = ↭-go (↭-prep x xs↭ys) rel
93134

94135
-- Skip reasoning about the first two elements
95136
step-swap : x y xs {ys zs : List A} (y ∷ x ∷ ys) IsRelatedTo zs
96137
xs ↭ ys (x ∷ y ∷ xs) IsRelatedTo zs
97-
step-swap x y xs rel xs↭ys = relTo (trans (swap x y xs↭ys) (begin rel))
138+
step-swap x y xs rel xs↭ys = ↭-go (↭-swap x y xs↭ys) rel
98139

99140
syntax step-prep x xs y↭z x↭y = x ∷ xs <⟨ x↭y ⟩ y↭z
100141
syntax step-swap x y xs y↭z x↭y = x ∷ y ∷ xs <<⟨ x↭y ⟩ y↭z

0 commit comments

Comments
 (0)