File tree Expand file tree Collapse file tree 2 files changed +26
-8
lines changed
src/Data/List/Relation/Binary/Permutation Expand file tree Collapse file tree 2 files changed +26
-8
lines changed Original file line number Diff line number Diff line change @@ -243,7 +243,12 @@ Additions to existing modules
243243 map-downFrom : ∀ (f : ℕ → A) n → map f (downFrom n) ≡ applyDownFrom f n
244244 ```
245245
246- * In ` Data.List.Relation.Binary.Permutation.PropositionalProperties ` :
246+ * In ` Data.List.Relation.Binary.Permutation.Propositional ` :
247+ ``` agda
248+ ↭⇒↭ₛ′ : IsEquivalence _≈_ → _↭_ ⇒ _↭ₛ′_
249+ ```
250+
251+ * In ` Data.List.Relation.Binary.Permutation.Propositional.Properties ` :
247252 ``` agda
248253 filter-↭ : ∀ (P? : Pred.Decidable P) → xs ↭ ys → filter P? xs ↭ filter P? ys
249254 ```
Original file line number Diff line number Diff line change @@ -88,19 +88,32 @@ data _↭_ : Rel (List A) a where
8888 }
8989
9090------------------------------------------------------------------------
91- -- _↭_ is equivalent to `Setoid`-based permutation
91+ -- _↭_ is finer than `Setoid`-based permutation for any equivalence on A
92+
93+ module _ {ℓ} {_≈_ : Rel A ℓ} (isEquivalence : IsEquivalence _≈_) where
94+
95+ private
96+ open module ↭ₛ′ = Permutation record { isEquivalence = isEquivalence }
97+ using ()
98+ renaming (_↭_ to _↭ₛ′_)
99+
100+ ↭⇒↭ₛ′ : _↭_ ⇒ _↭ₛ′_
101+ ↭⇒↭ₛ′ refl = ↭ₛ′.↭-refl
102+ ↭⇒↭ₛ′ (prep x p) = ↭ₛ′.↭-prep x (↭⇒↭ₛ′ p)
103+ ↭⇒↭ₛ′ (swap x y p) = ↭ₛ′.↭-swap x y (↭⇒↭ₛ′ p)
104+ ↭⇒↭ₛ′ (trans p q) = ↭ₛ′.↭-trans′ (↭⇒↭ₛ′ p) (↭⇒↭ₛ′ q)
105+
106+
107+ ------------------------------------------------------------------------
108+ -- _↭_ is equivalent to `Setoid`-based permutation on `≡.setoid A`
92109
93110private
94111 open module ↭ₛ = Permutation (≡.setoid A)
95112 using ()
96113 renaming (_↭_ to _↭ₛ_)
97114
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-
115+ ↭⇒↭ₛ : _↭_ ⇒ _↭ₛ_
116+ ↭⇒↭ₛ = ↭⇒↭ₛ′ ≡.isEquivalence
104117
105118↭ₛ⇒↭ : _↭ₛ_ ⇒ _↭_
106119↭ₛ⇒↭ (↭ₛ.refl xs≋ys) = ↭-reflexive-≋ xs≋ys
You can’t perform that action at this time.
0 commit comments