|
| 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