@@ -27,7 +27,7 @@ open import Level using (Level)
27
27
open import Relation.Nullary using (¬_; does; yes; no)
28
28
open import Relation.Nullary.Negation using (contradiction)
29
29
open import Relation.Unary using (Pred; Decidable) renaming (_⊆_ to _⋐_)
30
- open import Relation.Binary.Core using (_⇒_)
30
+ open import Relation.Binary.Core using (_⇒_; _Preserves_⟶_ )
31
31
open import Relation.Binary.Definitions
32
32
using (Reflexive; Transitive; _Respectsʳ_; _Respectsˡ_; _Respects_)
33
33
open import Relation.Binary.Bundles using (Setoid; Preorder)
@@ -39,7 +39,7 @@ open Setoid using (Carrier)
39
39
40
40
private
41
41
variable
42
- a p q ℓ : Level
42
+ a b p q r ℓ : Level
43
43
44
44
------------------------------------------------------------------------
45
45
-- Relational properties with _≋_ (pointwise equality)
@@ -210,6 +210,27 @@ module _ (S : Setoid a ℓ) where
210
210
++⁺ : ∀ {ws xs ys zs} → ws ⊆ xs → ys ⊆ zs → ws ++ ys ⊆ xs ++ zs
211
211
++⁺ ws⊆xs ys⊆zs = ⊆-trans S (++⁺ˡ _ ws⊆xs) (++⁺ʳ _ ys⊆zs)
212
212
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
+
213
234
------------------------------------------------------------------------
214
235
-- reverse
215
236
@@ -242,7 +263,6 @@ module _ (S : Setoid a ℓ) where
242
263
bs ∎
243
264
where open ⊆-Reasoning S
244
265
245
-
246
266
------------------------------------------------------------------------
247
267
-- filter
248
268
0 commit comments