Skip to content

Commit 1ddd742

Browse files
authored
Merge branch 'master' into refactor-Fin-properties
2 parents d7cffe0 + 29d1632 commit 1ddd742

File tree

11 files changed

+234
-61
lines changed

11 files changed

+234
-61
lines changed

CHANGELOG.md

Lines changed: 44 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -239,6 +239,18 @@ Additions to existing modules
239239
∙-cong-∣ : x ∣ y → a ∣ b → x ∙ a ∣ y ∙ b
240240
```
241241

242+
* In `Data.Fin.Base`:
243+
```agda
244+
_≰_ : Rel (Fin n) 0ℓ
245+
_≮_ : Rel (Fin n) 0ℓ
246+
```
247+
248+
* In `Data.Fin.Permutation`:
249+
```agda
250+
cast-id : .(m ≡ n) → Permutation m n
251+
swap : Permutation m n → Permutation (suc (suc m)) (suc (suc n))
252+
```
253+
242254
* In `Data.Fin.Permutation.Components`:
243255
```agda
244256
transpose-iij : (i j : Fin n) → transpose i i j ≡ j
@@ -249,6 +261,7 @@ Additions to existing modules
249261

250262
* In `Data.Fin.Properties`:
251263
```agda
264+
cast-involutive : .(eq₁ : m ≡ n) .(eq₂ : n ≡ m) → ∀ k → cast eq₁ (cast eq₂ k) ≡ k
252265
≡-irrelevant : Irrelevant {A = Fin n} _≡_
253266
≟-diag : (eq : i ≡ j) → (i ≟ j) ≡ yes eq
254267
≟-diag-refl : (i : Fin n) → (i ≟ i) ≡ yes refl
@@ -290,14 +303,31 @@ Additions to existing modules
290303
map-downFrom : ∀ (f : ℕ → A) n → map f (downFrom n) ≡ applyDownFrom f n
291304
```
292305

306+
* In `Data.List.Relation.Binary.Permutation.Homogeneous`:
307+
```agda
308+
onIndices : Permutation R xs ys → Fin.Permutation (length xs) (length ys)
309+
```
310+
293311
* In `Data.List.Relation.Binary.Permutation.Propositional`:
294312
```agda
295313
↭⇒↭ₛ′ : IsEquivalence _≈_ → _↭_ ⇒ _↭ₛ′_
296314
```
297315

316+
* In `Data.List.Relation.Binary.Permutation.Setoid.Properties`:
317+
```agda
318+
xs↭ys⇒|xs|≡|ys| : xs ↭ ys → length xs ≡ length ys
319+
¬x∷xs↭[] : ¬ (x ∷ xs ↭ [])
320+
onIndices-lookup : ∀ i → lookup xs i ≈ lookup ys (Inverse.to (onIndices xs↭ys) i)
321+
```
322+
298323
* In `Data.List.Relation.Binary.Permutation.Propositional.Properties`:
299324
```agda
300325
filter-↭ : ∀ (P? : Pred.Decidable P) → xs ↭ ys → filter P? xs ↭ filter P? ys
326+
```
327+
328+
* In `Data.List.Relation.Binary.Pointwise.Properties`:
329+
```agda
330+
lookup-cast : Pointwise R xs ys → .(∣xs∣≡∣ys∣ : length xs ≡ length ys) → ∀ i → R (lookup xs i) (lookup ys (cast ∣xs∣≡∣ys∣ i))
301331
```
302332

303333
* In `Data.List.NonEmpty.Properties`:
@@ -364,6 +394,20 @@ Additions to existing modules
364394
⊥-dec : Dec {a} ⊥
365395
```
366396

397+
* In `Relation.Unary`:
398+
```agda
399+
_⊥_ _⊥′_ : Pred A ℓ₁ → Pred A ℓ₂ → Set _
400+
```
401+
402+
* In `Relation.Unary.Properties`:
403+
```agda
404+
≬-symmetric : Sym _≬_ _≬_
405+
⊥-symmetric : Sym _⊥_ _⊥_
406+
≬-sym : Symmetric _≬_
407+
⊥-sym : Symmetric _⊥_
408+
≬⇒¬⊥ : _≬_ ⇒ (¬_ ∘₂ _⊥_)
409+
⊥⇒¬≬ : _⊥_ ⇒ (¬_ ∘₂ _≬_)
410+
367411
* In `Relation.Nullary.Negation.Core`:
368412
```agda
369413
contra-diagonal : (A → ¬ A) → ¬ A

src/Data/Fin/Base.agda

Lines changed: 7 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,7 @@ open import Level using (0ℓ)
1919
open import Relation.Binary.Core using (Rel)
2020
open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; refl; cong)
2121
open import Relation.Binary.Indexed.Heterogeneous.Core using (IRel)
22-
open import Relation.Nullary.Negation.Core using (contradiction)
22+
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
2323

2424
private
2525
variable
@@ -271,7 +271,7 @@ pinch {suc n} (suc i) (suc j) = suc (pinch i j)
271271
------------------------------------------------------------------------
272272
-- Order relations
273273

274-
infix 4 _≤_ _≥_ _<_ _>_
274+
infix 4 _≤_ _≥_ _<_ _>_ _≰_ _≮_
275275

276276
_≤_ : IRel Fin 0ℓ
277277
i ≤ j = toℕ i ℕ.≤ toℕ j
@@ -285,6 +285,11 @@ i < j = toℕ i ℕ.< toℕ j
285285
_>_ : IRel Fin 0ℓ
286286
i > j = toℕ i ℕ.> toℕ j
287287

288+
_≰_ : {n} Rel (Fin n) 0ℓ
289+
i ≰ j = ¬ (i ≤ j)
290+
291+
_≮_ : {n} Rel (Fin n) 0ℓ
292+
i ≮ j = ¬ (i < j)
288293

289294
------------------------------------------------------------------------
290295
-- An ordering view.

src/Data/Fin/Permutation.agda

Lines changed: 76 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -9,13 +9,20 @@
99
module Data.Fin.Permutation where
1010

1111
open import Data.Bool.Base using (true; false)
12+
<<<<<<< refactor-Fin-properties
1213
open import Data.Fin.Base using (Fin; suc; opposite; punchIn; punchOut)
1314
open import Data.Fin.Patterns using (0F)
1415
open import Data.Fin.Properties
1516
using (¬Fin0; _≟_; ≟-diag-refl; ≟-off-diag
1617
; opposite-involutive
1718
; punchInᵢ≢i; punchOut-punchIn; punchIn-punchOut
1819
; punchOut-cong; punchOut-cong′)
20+
=======
21+
open import Data.Fin.Base using (Fin; suc; cast; opposite; punchIn; punchOut)
22+
open import Data.Fin.Patterns using (0F; 1F)
23+
open import Data.Fin.Properties using (punchInᵢ≢i; punchOut-punchIn;
24+
punchOut-cong; punchOut-cong′; punchIn-punchOut; _≟_; ¬Fin0; cast-involutive)
25+
>>>>>>> master
1926
import Data.Fin.Permutation.Components as PC
2027
open import Data.Nat.Base using (ℕ; suc; zero)
2128
open import Data.Product.Base using (_,_)
@@ -25,7 +32,7 @@ open import Function.Construct.Identity using (↔-id)
2532
open import Function.Construct.Symmetry using (↔-sym)
2633
open import Function.Definitions using (StrictlyInverseˡ; StrictlyInverseʳ)
2734
open import Function.Properties.Inverse using (↔⇒↣)
28-
open import Function.Base using (_∘_)
35+
open import Function.Base using (_∘_; _∘′_)
2936
open import Level using (0ℓ)
3037
open import Relation.Binary.Core using (Rel)
3138
open import Relation.Nullary.Decidable.Core using (does; yes; no)
@@ -59,11 +66,15 @@ Permutation′ n = Permutation n n
5966
------------------------------------------------------------------------
6067
-- Helper functions
6168

62-
permutation : (f : Fin m Fin n) (g : Fin n Fin m)
63-
StrictlyInverseˡ _≡_ f g StrictlyInverseʳ _≡_ f g Permutation m n
69+
permutation : (f : Fin m Fin n)
70+
(g : Fin n Fin m)
71+
StrictlyInverseˡ _≡_ f g
72+
StrictlyInverseʳ _≡_ f g
73+
Permutation m n
6474
permutation = mk↔ₛ′
6575

6676
infixl 5 _⟨$⟩ʳ_ _⟨$⟩ˡ_
77+
6778
_⟨$⟩ʳ_ : Permutation m n Fin m Fin n
6879
_⟨$⟩ʳ_ = Inverse.to
6980

@@ -77,44 +88,66 @@ inverseʳ : ∀ (π : Permutation m n) {i} → π ⟨$⟩ʳ (π ⟨$⟩ˡ i) ≡
7788
inverseʳ π = Inverse.inverseˡ π refl
7889

7990
------------------------------------------------------------------------
80-
-- Equality
91+
-- Equality over permutations
8192

8293
infix 6 _≈_
94+
8395
_≈_ : Rel (Permutation m n) 0ℓ
8496
π ≈ ρ = i π ⟨$⟩ʳ i ≡ ρ ⟨$⟩ʳ i
8597

8698
------------------------------------------------------------------------
87-
-- Example permutations
99+
-- Permutation properties
88100

89-
-- Identity
90-
91-
id : Permutation′ n
101+
id : Permutation n n
92102
id = ↔-id _
93103

94-
-- Transpose two indices
95-
96-
transpose : Fin n Fin n Permutation′ n
97-
transpose i j = permutation (PC.transpose i j) (PC.transpose j i) (λ _ PC.transpose-inverse _ _) (λ _ PC.transpose-inverse _ _)
104+
flip : Permutation m n Permutation n m
105+
flip = ↔-sym
98106

99-
-- Reverse the order of indices
107+
infixr 9 _∘ₚ_
100108

109+
<<<<<<< refactor-Fin-properties
101110
reverse : Permutation′ n
102111
reverse = permutation opposite opposite opposite-involutive opposite-involutive
112+
=======
113+
_∘ₚ_ : Permutation m n Permutation n o Permutation m o
114+
π₁ ∘ₚ π₂ = π₂ ↔-∘ π₁
115+
>>>>>>> master
103116

104117
------------------------------------------------------------------------
105-
-- Operations
118+
-- Non-trivial identity
106119

107-
-- Composition
120+
cast-id : .(m ≡ n) Permutation m n
121+
cast-id m≡n = permutation
122+
(cast m≡n)
123+
(cast (sym m≡n))
124+
(cast-involutive m≡n (sym m≡n))
125+
(cast-involutive (sym m≡n) m≡n)
108126

109-
infixr 9 _∘ₚ_
110-
_∘ₚ_ : Permutation m n Permutation n o Permutation m o
111-
π₁ ∘ₚ π₂ = π₂ ↔-∘ π₁
127+
------------------------------------------------------------------------
128+
-- Transposition
112129

113-
-- Flip
130+
-- Transposes two elements in the permutation, keeping the remainder
131+
-- of the permutation the same
132+
transpose : Fin n Fin n Permutation n n
133+
transpose i j = permutation
134+
(PC.transpose i j)
135+
(PC.transpose j i)
136+
(λ _ PC.transpose-inverse _ _)
137+
(λ _ PC.transpose-inverse _ _)
114138

115-
flip : Permutation m n Permutation n m
116-
flip = ↔-sym
139+
------------------------------------------------------------------------
140+
-- Reverse
117141

142+
-- Reverses a permutation
143+
reverse : Permutation n n
144+
reverse = permutation
145+
opposite
146+
opposite
147+
PC.reverse-involutive
148+
PC.reverse-involutive
149+
150+
------------------------------------------------------------------------
118151
-- Element removal
119152
--
120153
-- `remove k [0 ↦ i₀, …, k ↦ iₖ, …, n ↦ iₙ]` yields
@@ -169,8 +202,14 @@ remove {m} {n} i π = permutation to from inverseˡ′ inverseʳ′
169202
≡⟨ punchOut-punchIn (πʳ i) ⟩
170203
j ∎
171204

172-
-- lift: takes a permutation m → n and creates a permutation (suc m) → (suc n)
205+
------------------------------------------------------------------------
206+
-- Lifting
207+
208+
-- Takes a permutation m → n and creates a permutation (suc m) → (suc n)
173209
-- by mapping 0 to 0 and applying the input permutation to everything else
210+
--
211+
-- Note: should be refactored as a special-case when we add the
212+
-- concatenation of two permutations
174213
lift₀ : Permutation m n Permutation (suc m) (suc n)
175214
lift₀ {m} {n} π = permutation to from inverseˡ′ inverseʳ′
176215
where
@@ -190,6 +229,9 @@ lift₀ {m} {n} π = permutation to from inverseˡ′ inverseʳ′
190229
inverseˡ′ 0F = refl
191230
inverseˡ′ (suc j) = cong suc (inverseʳ π)
192231

232+
------------------------------------------------------------------------
233+
-- Insertion
234+
193235
-- insert i j π is the permutation that maps i to j and otherwise looks like π
194236
-- it's roughly an inverse of remove
195237
insert : {m n} Fin (suc m) Fin (suc n) Permutation m n Permutation (suc m) (suc n)
@@ -239,6 +281,18 @@ insert {m} {n} i j π = permutation to from inverseˡ′ inverseʳ′
239281
≡⟨ punchIn-punchOut j≢k ⟩
240282
k ∎
241283

284+
------------------------------------------------------------------------
285+
-- Swapping
286+
287+
-- Takes a permutation m → n and creates a permutation
288+
-- suc (suc m) → suc (suc n) by mapping 0 to 1 and 1 to 0 and
289+
-- then applying the input permutation to everything else
290+
--
291+
-- Note: should be refactored as a special-case when we add the
292+
-- concatenation of two permutations
293+
swap : Permutation m n Permutation (suc (suc m)) (suc (suc n))
294+
swap π = transpose 0F 1F ∘ₚ lift₀ (lift₀ π)
295+
242296
------------------------------------------------------------------------
243297
-- Other properties
244298

src/Data/Fin/Properties.agda

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -294,6 +294,10 @@ cast-trans {m = suc _} {n = suc _} {o = suc _} eq₁ eq₂ zero = refl
294294
cast-trans {m = suc _} {n = suc _} {o = suc _} eq₁ eq₂ (suc k) =
295295
cong suc (cast-trans (ℕ.suc-injective eq₁) (ℕ.suc-injective eq₂) k)
296296

297+
cast-involutive : .(eq₁ : m ≡ n) .(eq₂ : n ≡ m)
298+
k cast eq₁ (cast eq₂ k) ≡ k
299+
cast-involutive eq₁ eq₂ k = trans (cast-trans eq₂ eq₁ k) (cast-is-id refl k)
300+
297301
------------------------------------------------------------------------
298302
-- Properties of _≤_
299303
------------------------------------------------------------------------

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

Lines changed: 28 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -8,12 +8,15 @@
88

99
module Data.List.Relation.Binary.Permutation.Homogeneous where
1010

11-
open import Data.List.Base using (List; _∷_)
11+
open import Data.List.Base using (List; _∷_; length)
1212
open import Data.List.Relation.Binary.Pointwise.Base as Pointwise
1313
using (Pointwise)
1414
import Data.List.Relation.Binary.Pointwise.Properties as Pointwise
1515
open import Data.Nat.Base using (ℕ; suc; _+_)
16+
open import Data.Fin.Base using (Fin; zero; suc; cast)
17+
import Data.Fin.Permutation as Fin
1618
open import Level using (Level; _⊔_)
19+
open import Function.Base using (_∘_)
1720
open import Relation.Binary.Core using (Rel; _⇒_)
1821
open import Relation.Binary.Bundles using (Setoid)
1922
open import Relation.Binary.Structures using (IsEquivalence)
@@ -23,6 +26,7 @@ private
2326
variable
2427
a r s : Level
2528
A : Set a
29+
R S : Rel A r
2630

2731
data Permutation {A : Set a} (R : Rel A r) : Rel (List A) (a ⊔ r) where
2832
refl : {xs ys} Pointwise R xs ys Permutation R xs ys
@@ -33,37 +37,39 @@ data Permutation {A : Set a} (R : Rel A r) : Rel (List A) (a ⊔ r) where
3337
------------------------------------------------------------------------
3438
-- The Permutation relation is an equivalence
3539

36-
module _ {R : Rel A r} where
40+
sym : Symmetric R Symmetric (Permutation R)
41+
sym R-sym (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys)
42+
sym R-sym (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (sym R-sym xs↭ys)
43+
sym R-sym (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (sym R-sym xs↭ys)
44+
sym R-sym (trans xs↭ys ys↭zs) = trans (sym R-sym ys↭zs) (sym R-sym xs↭ys)
3745

38-
sym : Symmetric R Symmetric (Permutation R)
39-
sym R-sym (refl xs∼ys) = refl (Pointwise.symmetric R-sym xs∼ys)
40-
sym R-sym (prep x∼x′ xs↭ys) = prep (R-sym x∼x′) (sym R-sym xs↭ys)
41-
sym R-sym (swap x∼x′ y∼y′ xs↭ys) = swap (R-sym y∼y′) (R-sym x∼x′) (sym R-sym xs↭ys)
42-
sym R-sym (trans xs↭ys ys↭zs) = trans (sym R-sym ys↭zs) (sym R-sym xs↭ys)
46+
isEquivalence : Reflexive R Symmetric R IsEquivalence (Permutation R)
47+
isEquivalence R-refl R-sym = record
48+
{ refl = refl (Pointwise.refl R-refl)
49+
; sym = sym R-sym
50+
; trans = trans
51+
}
4352

44-
isEquivalence : Reflexive R Symmetric R IsEquivalence (Permutation R)
45-
isEquivalence R-refl R-sym = record
46-
{ refl = refl (Pointwise.refl R-refl)
47-
; sym = sym R-sym
48-
; trans = trans
49-
}
53+
setoid : Reflexive R Symmetric R Setoid _ _
54+
setoid {R = R} R-refl R-sym = record
55+
{ isEquivalence = isEquivalence {R = R} R-refl R-sym
56+
}
5057

51-
setoid : Reflexive R Symmetric R Setoid _ _
52-
setoid R-refl R-sym = record
53-
{ isEquivalence = isEquivalence R-refl R-sym
54-
}
55-
56-
map : {R : Rel A r} {S : Rel A s}
57-
(R ⇒ S) (Permutation R ⇒ Permutation S)
58+
map : (R ⇒ S) (Permutation R ⇒ Permutation S)
5859
map R⇒S (refl xs∼ys) = refl (Pointwise.map R⇒S xs∼ys)
5960
map R⇒S (prep e xs∼ys) = prep (R⇒S e) (map R⇒S xs∼ys)
6061
map R⇒S (swap e₁ e₂ xs∼ys) = swap (R⇒S e₁) (R⇒S e₂) (map R⇒S xs∼ys)
6162
map R⇒S (trans xs∼ys ys∼zs) = trans (map R⇒S xs∼ys) (map R⇒S ys∼zs)
6263

6364
-- Measures the number of constructors, can be useful for termination proofs
64-
65-
steps : {R : Rel A r} {xs ys} Permutation R xs ys
65+
steps : {xs ys} Permutation R xs ys
6666
steps (refl _) = 1
6767
steps (prep _ xs↭ys) = suc (steps xs↭ys)
6868
steps (swap _ _ xs↭ys) = suc (steps xs↭ys)
6969
steps (trans xs↭ys ys↭zs) = steps xs↭ys + steps ys↭zs
70+
71+
onIndices : {xs ys} Permutation R xs ys Fin.Permutation (length xs) (length ys)
72+
onIndices (refl ≋) = Fin.cast-id (Pointwise.Pointwise-length ≋)
73+
onIndices (prep e xs↭ys) = Fin.lift₀ (onIndices xs↭ys)
74+
onIndices (swap e f xs↭ys) = Fin.swap (onIndices xs↭ys)
75+
onIndices (trans ↭₁ ↭₂) = onIndices ↭₁ Fin.∘ₚ onIndices ↭₂

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

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -31,7 +31,7 @@ open ≋ S using (_≋_; _∷_; ≋-refl; ≋-sym; ≋-trans)
3131
-- Definition, based on `Homogeneous`
3232

3333
open Homogeneous public
34-
using (refl; prep; swap; trans)
34+
using (refl; prep; swap; trans; onIndices)
3535

3636
infix 3 _↭_
3737

0 commit comments

Comments
 (0)