Skip to content

Commit b0ad861

Browse files
authored
fixes #2375 (#2377)
* fixes #2375 * removed redundant `Membership` instances * fix merge conflict error
1 parent b671f95 commit b0ad861

File tree

2 files changed

+25
-3
lines changed

2 files changed

+25
-3
lines changed

CHANGELOG.md

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -382,6 +382,8 @@ Additions to existing modules
382382

383383
* In `Data.List.Relation.Binary.Subset.Setoid.Properties`:
384384
```agda
385+
map⁺ : f Preserves _≈_ ⟶ _≈′_ → as ⊆ bs → map f as ⊆′ map f bs
386+
385387
reverse-selfAdjoint : as ⊆ reverse bs → reverse as ⊆ bs
386388
reverse⁺ : as ⊆ bs → reverse as ⊆ reverse bs
387389
reverse⁻ : reverse as ⊆ reverse bs → as ⊆ bs

src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda

Lines changed: 23 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,7 @@ open import Level using (Level)
2727
open import Relation.Nullary using (¬_; does; yes; no)
2828
open import Relation.Nullary.Negation using (contradiction)
2929
open import Relation.Unary using (Pred; Decidable) renaming (_⊆_ to _⋐_)
30-
open import Relation.Binary.Core using (_⇒_)
30+
open import Relation.Binary.Core using (_⇒_; _Preserves_⟶_)
3131
open import Relation.Binary.Definitions
3232
using (Reflexive; Transitive; _Respectsʳ_; _Respectsˡ_; _Respects_)
3333
open import Relation.Binary.Bundles using (Setoid; Preorder)
@@ -39,7 +39,7 @@ open Setoid using (Carrier)
3939

4040
private
4141
variable
42-
a p q ℓ : Level
42+
a b p q r: Level
4343

4444
------------------------------------------------------------------------
4545
-- Relational properties with _≋_ (pointwise equality)
@@ -210,6 +210,27 @@ module _ (S : Setoid a ℓ) where
210210
++⁺ : {ws xs ys zs} ws ⊆ xs ys ⊆ zs ws ++ ys ⊆ xs ++ zs
211211
++⁺ ws⊆xs ys⊆zs = ⊆-trans S (++⁺ˡ _ ws⊆xs) (++⁺ʳ _ ys⊆zs)
212212

213+
------------------------------------------------------------------------
214+
-- map
215+
216+
module _ (S : Setoid a ℓ) (R : Setoid b r) where
217+
218+
private
219+
module S = Setoid S
220+
module R = Setoid R
221+
222+
module S⊆ = Subset S
223+
module R⊆ = Subset R
224+
225+
open Membershipₚ
226+
227+
map⁺ : {as bs} {f : S.Carrier R.Carrier}
228+
f Preserves S._≈_ ⟶ R._≈_
229+
as S⊆.⊆ bs map f as R⊆.⊆ map f bs
230+
map⁺ {f = f} f-pres as⊆bs v∈f[as] =
231+
let x , x∈as , v≈f[x] = ∈-map⁻ S R v∈f[as] in
232+
∈-resp-≈ R (R.sym v≈f[x]) (∈-map⁺ S R f-pres (as⊆bs x∈as))
233+
213234
------------------------------------------------------------------------
214235
-- reverse
215236

@@ -242,7 +263,6 @@ module _ (S : Setoid a ℓ) where
242263
bs ∎
243264
where open ⊆-Reasoning S
244265

245-
246266
------------------------------------------------------------------------
247267
-- filter
248268

0 commit comments

Comments
 (0)