Skip to content

Commit 995fdab

Browse files
authored
Tidy up functional vector permutation #2066 (#2312)
* added structure and bundle; tidied up * added structure and bundle; tidied up * fix order of entries * blank line * standardise `variable` names * alignment
1 parent cb4d3ac commit 995fdab

File tree

3 files changed

+53
-27
lines changed

3 files changed

+53
-27
lines changed

CHANGELOG.md

+19-17
Original file line numberDiff line numberDiff line change
@@ -79,7 +79,22 @@ New modules
7979
```agda
8080
Algebra.Module.Construct.Idealization
8181
```
82-
82+
83+
* `Data.Vec.Functional.Relation.Binary.Permutation`, defining:
84+
```agda
85+
_↭_ : IRel (Vector A) _
86+
```
87+
88+
* `Data.Vec.Functional.Relation.Binary.Permutation.Properties` of the above:
89+
```agda
90+
↭-refl : Reflexive (Vector A) _↭_
91+
↭-reflexive : xs ≡ ys → xs ↭ ys
92+
↭-sym : Symmetric (Vector A) _↭_
93+
↭-trans : Transitive (Vector A) _↭_
94+
isIndexedEquivalence : {A : Set a} → IsIndexedEquivalence (Vector A) _↭_
95+
indexedSetoid : {A : Set a} → IndexedSetoid ℕ a _
96+
```
97+
8398
* `Function.Relation.Binary.Equality`
8499
```agda
85100
setoid : Setoid a₁ a₂ → Setoid b₁ b₂ → Setoid _ _
@@ -358,26 +373,13 @@ Additions to existing modules
358373
WeaklyDecidable : Set _
359374
```
360375

361-
* Added new definitions in `Relation.Unary`
362-
```
363-
Stable : Pred A ℓ → Set _
364-
WeaklyDecidable : Pred A ℓ → Set _
365-
```
366-
367376
* Added new proof in `Relation.Nullary.Decidable`:
368377
```agda
369378
⌊⌋-map′ : (a? : Dec A) → ⌊ map′ t f a? ⌋ ≡ ⌊ a? ⌋
370379
```
371380

372-
* Added module `Data.Vec.Functional.Relation.Binary.Permutation`:
373-
```agda
374-
_↭_ : IRel (Vector A) _
381+
* Added new definitions in `Relation.Unary`
375382
```
376-
377-
* Added new file `Data.Vec.Functional.Relation.Binary.Permutation.Properties`:
378-
```agda
379-
↭-refl : Reflexive (Vector A) _↭_
380-
↭-reflexive : xs ≡ ys → xs ↭ ys
381-
↭-sym : Symmetric (Vector A) _↭_
382-
↭-trans : Transitive (Vector A) _↭_
383+
Stable : Pred A ℓ → Set _
384+
WeaklyDecidable : Pred A ℓ → Set _
383385
```

src/Data/Vec/Functional/Relation/Binary/Permutation.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -17,8 +17,8 @@ open import Relation.Binary.PropositionalEquality.Core using (_≡_)
1717

1818
private
1919
variable
20-
: Level
21-
A : Set
20+
a : Level
21+
A : Set a
2222

2323
infix 3 _↭_
2424

src/Data/Vec/Functional/Relation/Binary/Permutation/Properties.agda

+32-8
Original file line numberDiff line numberDiff line change
@@ -15,18 +15,20 @@ open import Data.Fin.Permutation using (id; flip; _⟨$⟩ʳ_; inverseʳ; _∘
1515
open import Data.Vec.Functional
1616
open import Data.Vec.Functional.Relation.Binary.Permutation
1717
open import Relation.Binary.PropositionalEquality
18-
using (refl; trans; _≡_; cong; module ≡-Reasoning)
19-
open import Relation.Binary.Indexed.Heterogeneous.Definitions
20-
21-
open ≡-Reasoning
18+
using (_≡_; refl; trans; cong; module ≡-Reasoning)
19+
open import Relation.Binary.Indexed.Heterogeneous
2220

2321
private
2422
variable
25-
: Level
26-
A : Set
23+
a : Level
24+
A : Set a
2725
n :
2826
xs ys : Vector A n
2927

28+
29+
------------------------------------------------------------------------
30+
-- Basics
31+
3032
↭-refl : Reflexive (Vector A) _↭_
3133
↭-refl = id , λ _ refl
3234

@@ -36,10 +38,32 @@ private
3638
↭-sym : Symmetric (Vector A) _↭_
3739
proj₁ (↭-sym (xs↭ys , _)) = flip xs↭ys
3840
proj₂ (↭-sym {x = xs} {ys} (xs↭ys , xs↭ys≡)) i = begin
39-
ys (flip xs↭ys ⟨$⟩ʳ i) ≡˘⟨ xs↭ys≡ _
41+
ys (flip xs↭ys ⟨$⟩ʳ i) ⟨ xs↭ys≡ _
4042
xs (xs↭ys ⟨$⟩ʳ (flip xs↭ys ⟨$⟩ʳ i)) ≡⟨ cong xs (inverseʳ xs↭ys) ⟩
41-
xs i ∎
43+
xs i ∎
44+
where open ≡-Reasoning
4245

4346
↭-trans : Transitive (Vector A) _↭_
4447
proj₁ (↭-trans (xs↭ys , _) (ys↭zs , _)) = ys↭zs ∘ₚ xs↭ys
4548
proj₂ (↭-trans (_ , xs↭ys) (_ , ys↭zs)) _ = trans (xs↭ys _) (ys↭zs _)
49+
50+
------------------------------------------------------------------------
51+
-- Structure
52+
53+
isIndexedEquivalence : IsIndexedEquivalence (Vector A) _↭_
54+
isIndexedEquivalence {A = A} = record
55+
{ refl = ↭-refl
56+
; sym = ↭-sym
57+
; trans = λ {n₁ n₂ n₃} {xs : Vector A n₁} {ys : Vector A n₂} {zs : Vector A n₃}
58+
xs↭ys ys↭zs ↭-trans {i = n₁} {i = xs} xs↭ys ys↭zs
59+
}
60+
61+
------------------------------------------------------------------------
62+
-- Bundle
63+
64+
indexedSetoid : {A : Set a} IndexedSetoid ℕ a _
65+
indexedSetoid {A = A} = record
66+
{ Carrier = Vector A
67+
; _≈_ = _↭_
68+
; isEquivalence = isIndexedEquivalence
69+
}

0 commit comments

Comments
 (0)