Skip to content

Commit 3146ae5

Browse files
guilhermehasTanebguialvares
authored andcommitted
Added functional vector permutation (#2066)
* added functional vector permutation * added one line to CHANGELOG * added permutation properties * Added Base to imports Co-authored-by: Nathan van Doorn <[email protected]> * Added Base to import Co-authored-by: Nathan van Doorn <[email protected]> * Added core to import Co-authored-by: Nathan van Doorn <[email protected]> * added definitions * removed unnecessary zs * renamed types in changelog * removed unnecessary code * added more to changelog * added end of changelog --------- Co-authored-by: Nathan van Doorn <[email protected]> Co-authored-by: Guilherme <[email protected]>
1 parent 68b17e0 commit 3146ae5

File tree

3 files changed

+84
-0
lines changed

3 files changed

+84
-0
lines changed

CHANGELOG.md

+13
Original file line numberDiff line numberDiff line change
@@ -187,3 +187,16 @@ Additions to existing modules
187187
```agda
188188
⌊⌋-map′ : (a? : Dec A) → ⌊ map′ t f a? ⌋ ≡ ⌊ a? ⌋
189189
```
190+
191+
* Added module `Data.Vec.Functional.Relation.Binary.Permutation`:
192+
```agda
193+
_↭_ : IRel (Vector A) _
194+
```
195+
196+
* Added new file `Data.Vec.Functional.Relation.Binary.Permutation.Properties`:
197+
```agda
198+
↭-refl : Reflexive (Vector A) _↭_
199+
↭-reflexive : xs ≡ ys → xs ↭ ys
200+
↭-sym : Symmetric (Vector A) _↭_
201+
↭-trans : Transitive (Vector A) _↭_
202+
```
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,26 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Permutation relations over Vector
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.Vec.Functional.Relation.Binary.Permutation where
10+
11+
open import Level using (Level)
12+
open import Data.Product.Base using (Σ-syntax)
13+
open import Data.Fin.Permutation using (Permutation; _⟨$⟩ʳ_)
14+
open import Data.Vec.Functional using (Vector)
15+
open import Relation.Binary.Indexed.Heterogeneous.Core using (IRel)
16+
open import Relation.Binary.PropositionalEquality.Core using (_≡_)
17+
18+
private
19+
variable
20+
: Level
21+
A : Set
22+
23+
infix 3 _↭_
24+
25+
_↭_ : IRel (Vector A) _
26+
xs ↭ ys = Σ[ ρ ∈ Permutation _ _ ] ( i xs (ρ ⟨$⟩ʳ i) ≡ ys i)
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,45 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Properties of permutation
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.Vec.Functional.Relation.Binary.Permutation.Properties where
10+
11+
open import Level using (Level)
12+
open import Data.Product.Base using (_,_; proj₁; proj₂)
13+
open import Data.Nat.Base using (ℕ)
14+
open import Data.Fin.Permutation using (id; flip; _⟨$⟩ʳ_; inverseʳ; _∘ₚ_)
15+
open import Data.Vec.Functional
16+
open import Data.Vec.Functional.Relation.Binary.Permutation
17+
open import Relation.Binary.PropositionalEquality
18+
using (refl; trans; _≡_; cong; module ≡-Reasoning)
19+
open import Relation.Binary.Indexed.Heterogeneous.Definitions
20+
21+
open ≡-Reasoning
22+
23+
private
24+
variable
25+
: Level
26+
A : Set
27+
n :
28+
xs ys : Vector A n
29+
30+
↭-refl : Reflexive (Vector A) _↭_
31+
↭-refl = id , λ _ refl
32+
33+
↭-reflexive : xs ≡ ys xs ↭ ys
34+
↭-reflexive refl = ↭-refl
35+
36+
↭-sym : Symmetric (Vector A) _↭_
37+
proj₁ (↭-sym (xs↭ys , _)) = flip xs↭ys
38+
proj₂ (↭-sym {x = xs} {ys} (xs↭ys , xs↭ys≡)) i = begin
39+
ys (flip xs↭ys ⟨$⟩ʳ i) ≡˘⟨ xs↭ys≡ _ ⟩
40+
xs (xs↭ys ⟨$⟩ʳ (flip xs↭ys ⟨$⟩ʳ i)) ≡⟨ cong xs (inverseʳ xs↭ys) ⟩
41+
xs i ∎
42+
43+
↭-trans : Transitive (Vector A) _↭_
44+
proj₁ (↭-trans (xs↭ys , _) (ys↭zs , _)) = ys↭zs ∘ₚ xs↭ys
45+
proj₂ (↭-trans (_ , xs↭ys) (_ , ys↭zs)) _ = trans (xs↭ys _) (ys↭zs _)

0 commit comments

Comments
 (0)