Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 43514e1

Browse files
committed
feat(data/option/n_ary): Binary map of options (#16763)
Define `option.map₂`, the binary map of options.
1 parent a276716 commit 43514e1

File tree

8 files changed

+243
-29
lines changed

8 files changed

+243
-29
lines changed

src/data/finset/functor.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yaël Dillies, Scott Morrison
55
-/
66
import data.finset.lattice
7+
import data.finset.n_ary
78
import data.multiset.functor
89

910
/-!
@@ -64,6 +65,12 @@ instance : applicative finset :=
6465
@[simp] lemma seq_left_def (s : finset α) (t : finset β) : s <* t = if t = ∅ thenelse s := rfl
6566
@[simp] lemma seq_right_def (s : finset α) (t : finset β) : s *> t = if s = ∅ thenelse t := rfl
6667

68+
/-- `finset.image₂` in terms of monadic operations. Note that this can't be taken as the definition
69+
because of the lack of universe polymorphism. -/
70+
lemma image₂_def {α β γ : Type*} (f : α → β → γ) (s : finset α) (t : finset β) :
71+
image₂ f s t = f <$> s <*> t :=
72+
by { ext, simp [mem_sup] }
73+
6774
instance : is_lawful_applicative finset :=
6875
{ seq_left_eq := λ α β s t, begin
6976
rw [seq_def, fmap_def, seq_left_def],

src/data/finset/n_ary.lean

Lines changed: 21 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -13,8 +13,8 @@ This file defines `finset.image₂`, the binary image of finsets. This is the fi
1313
1414
## Notes
1515
16-
This file is very similar to the n-ary section of `data.set.basic` and to `order.filter.n_ary`.
17-
Please keep them in sync.
16+
This file is very similar to the n-ary section of `data.set.basic`, to `order.filter.n_ary` and to
17+
`data.option.n_ary`. Please keep them in sync.
1818
1919
We do not define `finset.image₃` as its only purpose would be to prove properties of `finset.image₂`
2020
and `set.image2` already fulfills this task.
@@ -206,6 +206,17 @@ lemma image₂_swap (f : α → β → γ) (s : finset α) (t : finset β) :
206206
image₂ f s t = image₂ (λ a b, f b a) t s :=
207207
coe_injective $ by { push_cast, exact image2_swap _ _ _ }
208208

209+
@[simp] lemma image₂_mk_eq_product [decidable_eq α] [decidable_eq β] (s : finset α) (t : finset β) :
210+
image₂ prod.mk s t = s ×ˢ t :=
211+
by ext; simp [prod.ext_iff]
212+
213+
@[simp] lemma image₂_curry (f : α × β → γ) (s : finset α) (t : finset β) :
214+
image₂ (curry f) s t = (s ×ˢ t).image f :=
215+
by { classical, rw [←image₂_mk_eq_product, image_image₂, curry] }
216+
217+
@[simp] lemma image_uncurry_product (f : α → β → γ) (s : finset α) (t : finset β) :
218+
(s ×ˢ t).image (uncurry f) = image₂ f s t := by rw [←image₂_curry, curry_uncurry]
219+
209220
@[simp] lemma image₂_left [decidable_eq α] (h : t.nonempty) : image₂ (λ x y, x) s t = s :=
210221
coe_injective $ by { push_cast, exact image2_left h }
211222

@@ -235,25 +246,25 @@ lemma image_image₂_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ :
235246
(image₂ f s t).image g = image₂ f' (s.image g₁) (t.image g₂) :=
236247
coe_injective $ by { push_cast, exact image_image2_distrib h_distrib }
237248

238-
/-- Symmetric of `finset.image₂_image_left_comm`. -/
249+
/-- Symmetric statement to `finset.image₂_image_left_comm`. -/
239250
lemma image_image₂_distrib_left {g : γ → δ} {f' : α' → β → δ} {g' : α → α'}
240251
(h_distrib : ∀ a b, g (f a b) = f' (g' a) b) :
241252
(image₂ f s t).image g = image₂ f' (s.image g') t :=
242253
coe_injective $ by { push_cast, exact image_image2_distrib_left h_distrib }
243254

244-
/-- Symmetric of `finset.image_image₂_right_comm`. -/
255+
/-- Symmetric statement to `finset.image_image₂_right_comm`. -/
245256
lemma image_image₂_distrib_right {g : γ → δ} {f' : α → β' → δ} {g' : β → β'}
246257
(h_distrib : ∀ a b, g (f a b) = f' a (g' b)) :
247258
(image₂ f s t).image g = image₂ f' s (t.image g') :=
248259
coe_injective $ by { push_cast, exact image_image2_distrib_right h_distrib }
249260

250-
/-- Symmetric of `finset.image_image₂_distrib_left`. -/
261+
/-- Symmetric statement to `finset.image_image₂_distrib_left`. -/
251262
lemma image₂_image_left_comm {f : α' → β → γ} {g : α → α'} {f' : α → β → δ} {g' : δ → γ}
252263
(h_left_comm : ∀ a b, f (g a) b = g' (f' a b)) :
253264
image₂ f (s.image g) t = (image₂ f' s t).image g' :=
254265
(image_image₂_distrib_left $ λ a b, (h_left_comm a b).symm).symm
255266

256-
/-- Symmetric of `finset.image_image₂_distrib_right`. -/
267+
/-- Symmetric statement to `finset.image_image₂_distrib_right`. -/
257268
lemma image_image₂_right_comm {f : α → β' → γ} {g : β → β'} {f' : α → β → δ} {g' : δ → γ}
258269
(h_right_comm : ∀ a b, f a (g b) = g' (f' a b)) :
259270
image₂ f s (t.image g) = (image₂ f' s t).image g' :=
@@ -278,25 +289,25 @@ lemma image_image₂_antidistrib {g : γ → δ} {f' : β' → α' → δ} {g₁
278289
(image₂ f s t).image g = image₂ f' (t.image g₁) (s.image g₂) :=
279290
by { rw image₂_swap f, exact image_image₂_distrib (λ _ _, h_antidistrib _ _) }
280291

281-
/-- Symmetric of `finset.image₂_image_left_anticomm`. -/
292+
/-- Symmetric statement to `finset.image₂_image_left_anticomm`. -/
282293
lemma image_image₂_antidistrib_left {g : γ → δ} {f' : β' → α → δ} {g' : β → β'}
283294
(h_antidistrib : ∀ a b, g (f a b) = f' (g' b) a) :
284295
(image₂ f s t).image g = image₂ f' (t.image g') s :=
285296
coe_injective $ by { push_cast, exact image_image2_antidistrib_left h_antidistrib }
286297

287-
/-- Symmetric of `finset.image_image₂_right_anticomm`. -/
298+
/-- Symmetric statement to `finset.image_image₂_right_anticomm`. -/
288299
lemma image_image₂_antidistrib_right {g : γ → δ} {f' : β → α' → δ} {g' : α → α'}
289300
(h_antidistrib : ∀ a b, g (f a b) = f' b (g' a)) :
290301
(image₂ f s t).image g = image₂ f' t (s.image g') :=
291302
coe_injective $ by { push_cast, exact image_image2_antidistrib_right h_antidistrib }
292303

293-
/-- Symmetric of `finset.image_image₂_antidistrib_left`. -/
304+
/-- Symmetric statement to `finset.image_image₂_antidistrib_left`. -/
294305
lemma image₂_image_left_anticomm {f : α' → β → γ} {g : α → α'} {f' : β → α → δ} {g' : δ → γ}
295306
(h_left_anticomm : ∀ a b, f (g a) b = g' (f' b a)) :
296307
image₂ f (s.image g) t = (image₂ f' t s).image g' :=
297308
(image_image₂_antidistrib_left $ λ a b, (h_left_anticomm b a).symm).symm
298309

299-
/-- Symmetric of `finset.image_image₂_antidistrib_right`. -/
310+
/-- Symmetric statement to `finset.image_image₂_antidistrib_right`. -/
300311
lemma image_image₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : β → α → δ} {g' : δ → γ}
301312
(h_right_anticomm : ∀ a b, f a (g b) = g' (f' b a)) :
302313
image₂ f s (t.image g) = (image₂ f' t s).image g' :=

src/data/option/basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -35,8 +35,10 @@ namespace option
3535
variables {α β γ δ : Type*}
3636

3737
lemma coe_def : (coe : α → option α) = some := rfl
38+
lemma some_eq_coe (a : α) : some a = a := rfl
3839

3940
lemma some_ne_none (x : α) : some x ≠ none := λ h, option.no_confusion h
41+
@[simp] lemma coe_ne_none (a : α) : (a : option α) ≠ none .
4042

4143
protected lemma «forall» {p : option α → Prop} : (∀ x, p x) ↔ p none ∧ ∀ x, p (some x) :=
4244
⟨λ h, ⟨h _, λ x, h _⟩, λ h x, option.cases_on x h.1 h.2

src/data/option/n_ary.lean

Lines changed: 170 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,170 @@
1+
/-
2+
Copyright (c) 2022 Yaël Dillies. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yaël Dillies
5+
-/
6+
import data.option.basic
7+
8+
/-!
9+
# Binary map of options
10+
11+
This file defines the binary map of `option`. This is mostly useful to define pointwise operations
12+
on intervals.
13+
14+
## Main declarations
15+
16+
* `option.map₂`: Binary map of options.
17+
18+
## Notes
19+
20+
This file is very similar to the n-ary section of `data.set.basic`, to `data.finset.n_ary` and to
21+
`order.filter.n_ary`. Please keep them in sync.
22+
23+
We do not define `option.map₃` as its only purpose so far would be to prove properties of
24+
`option.map₂` and casing already fulfills this task.
25+
-/
26+
27+
open function
28+
29+
namespace option
30+
variables {α α' β β' γ γ' δ δ' ε ε' : Type*} {f : α → β → γ} {a : option α} {b : option β}
31+
{c : option γ}
32+
33+
/-- The image of a binary function `f : α → β → γ` as a function `option α → option β → option γ`.
34+
Mathematically this should be thought of as the image of the corresponding function `α × β → γ`. -/
35+
def map₂ (f : α → β → γ) (a : option α) (b : option β) : option γ := a.bind $ λ a, b.map $ f a
36+
37+
/-- `option.map₂` in terms of monadic operations. Note that this can't be taken as the definition
38+
because of the lack of universe polymorphism. -/
39+
lemma map₂_def {α β γ : Type*} (f : α → β → γ) (a : option α) (b : option β) :
40+
map₂ f a b = f <$> a <*> b := by cases a; refl
41+
42+
@[simp] lemma map₂_some_some (f : α → β → γ) (a : α) (b : β) : map₂ f (some a) (some b) = f a b :=
43+
rfl
44+
lemma map₂_coe_coe (f : α → β → γ) (a : α) (b : β) : map₂ f a b = f a b := rfl
45+
@[simp] lemma map₂_none_left (f : α → β → γ) (b : option β) : map₂ f none b = none := rfl
46+
@[simp] lemma map₂_none_right (f : α → β → γ) (a : option α) : map₂ f a none = none :=
47+
by cases a; refl
48+
@[simp] lemma map₂_coe_left (f : α → β → γ) (a : α) (b : option β) :
49+
map₂ f a b = b.map (λ b, f a b) := rfl
50+
@[simp] lemma map₂_coe_right (f : α → β → γ) (a : option α) (b : β) :
51+
map₂ f a b = a.map (λ a, f a b) := rfl
52+
53+
@[simp] lemma mem_map₂_iff {c : γ} : c ∈ map₂ f a b ↔ ∃ a' b', a' ∈ a ∧ b' ∈ b ∧ f a' b' = c :=
54+
by simp [map₂]
55+
56+
@[simp] lemma map₂_eq_none_iff : map₂ f a b = none ↔ a = none ∨ b = none :=
57+
by cases a; cases b; simp
58+
59+
lemma map₂_swap (f : α → β → γ) (a : option α) (b : option β) :
60+
map₂ f a b = map₂ (λ a b, f b a) b a :=
61+
by cases a; cases b; refl
62+
63+
lemma map_map₂ (f : α → β → γ) (g : γ → δ) : (map₂ f a b).map g = map₂ (λ a b, g (f a b)) a b :=
64+
by cases a; cases b; refl
65+
66+
lemma map₂_map_left (f : γ → β → δ) (g : α → γ) :
67+
map₂ f (a.map g) b = map₂ (λ a b, f (g a) b) a b :=
68+
by cases a; refl
69+
70+
lemma map₂_map_right (f : α → γ → δ) (g : β → γ) :
71+
map₂ f a (b.map g) = map₂ (λ a b, f a (g b)) a b :=
72+
by cases b; refl
73+
74+
@[simp] lemma map₂_curry (f : α × β → γ) (a : option α) (b : option β) :
75+
map₂ (curry f) a b = option.map f (map₂ prod.mk a b) := (map_map₂ _ _).symm
76+
77+
@[simp] lemma map_uncurry (f : α → β → γ) (x : option (α × β)) :
78+
x.map (uncurry f) = map₂ f (x.map prod.fst) (x.map prod.snd) := by cases x; refl
79+
80+
/-!
81+
### Algebraic replacement rules
82+
83+
A collection of lemmas to transfer associativity, commutativity, distributivity, ... of operations
84+
to the associativity, commutativity, distributivity, ... of `option.map₂` of those operations.
85+
The proof pattern is `map₂_lemma operation_lemma`. For example, `map₂_comm mul_comm` proves that
86+
`map₂ (*) a b = map₂ (*) g f` in a `comm_semigroup`.
87+
-/
88+
89+
lemma map₂_assoc {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} {g' : β → γ → ε'}
90+
(h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) :
91+
map₂ f (map₂ g a b) c = map₂ f' a (map₂ g' b c) :=
92+
by cases a; cases b; cases c; simp [h_assoc]
93+
94+
lemma map₂_comm {g : β → α → γ} (h_comm : ∀ a b, f a b = g b a) : map₂ f a b = map₂ g b a :=
95+
by cases a; cases b; simp [h_comm]
96+
97+
lemma map₂_left_comm {f : α → δ → ε} {g : β → γ → δ} {f' : α → γ → δ'} {g' : β → δ' → ε}
98+
(h_left_comm : ∀ a b c, f a (g b c) = g' b (f' a c)) :
99+
map₂ f a (map₂ g b c) = map₂ g' b (map₂ f' a c) :=
100+
by cases a; cases b; cases c; simp [h_left_comm]
101+
102+
lemma map₂_right_comm {f : δ → γ → ε} {g : α → β → δ} {f' : α → γ → δ'} {g' : δ' → β → ε}
103+
(h_right_comm : ∀ a b c, f (g a b) c = g' (f' a c) b) :
104+
map₂ f (map₂ g a b) c = map₂ g' (map₂ f' a c) b :=
105+
by cases a; cases b; cases c; simp [h_right_comm]
106+
107+
lemma map_map₂_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ : α → α'} {g₂ : β → β'}
108+
(h_distrib : ∀ a b, g (f a b) = f' (g₁ a) (g₂ b)) :
109+
(map₂ f a b).map g = map₂ f' (a.map g₁) (b.map g₂) :=
110+
by cases a; cases b; simp [h_distrib]
111+
112+
/-!
113+
The following symmetric restatement are needed because unification has a hard time figuring all the
114+
functions if you symmetrize on the spot. This is also how the other n-ary APIs do it.
115+
-/
116+
117+
/-- Symmetric statement to `option.map₂_map_left_comm`. -/
118+
lemma map_map₂_distrib_left {g : γ → δ} {f' : α' → β → δ} {g' : α → α'}
119+
(h_distrib : ∀ a b, g (f a b) = f' (g' a) b) :
120+
(map₂ f a b).map g = map₂ f' (a.map g') b :=
121+
by cases a; cases b; simp [h_distrib]
122+
123+
/-- Symmetric statement to `option.map_map₂_right_comm`. -/
124+
lemma map_map₂_distrib_right {g : γ → δ} {f' : α → β' → δ} {g' : β → β'}
125+
(h_distrib : ∀ a b, g (f a b) = f' a (g' b)) :
126+
(map₂ f a b).map g = map₂ f' a (b.map g') :=
127+
by cases a; cases b; simp [h_distrib]
128+
129+
/-- Symmetric statement to `option.map_map₂_distrib_left`. -/
130+
lemma map₂_map_left_comm {f : α' → β → γ} {g : α → α'} {f' : α → β → δ} {g' : δ → γ}
131+
(h_left_comm : ∀ a b, f (g a) b = g' (f' a b)) :
132+
map₂ f (a.map g) b = (map₂ f' a b).map g' :=
133+
by cases a; cases b; simp [h_left_comm]
134+
135+
/-- Symmetric statement to `option.map_map₂_distrib_right`. -/
136+
lemma map_map₂_right_comm {f : α → β' → γ} {g : β → β'} {f' : α → β → δ} {g' : δ → γ}
137+
(h_right_comm : ∀ a b, f a (g b) = g' (f' a b)) :
138+
map₂ f a (b.map g) = (map₂ f' a b).map g' :=
139+
by cases a; cases b; simp [h_right_comm]
140+
141+
lemma map_map₂_antidistrib {g : γ → δ} {f' : β' → α' → δ} {g₁ : β → β'} {g₂ : α → α'}
142+
(h_antidistrib : ∀ a b, g (f a b) = f' (g₁ b) (g₂ a)) :
143+
(map₂ f a b).map g = map₂ f' (b.map g₁) (a.map g₂) :=
144+
by cases a; cases b; simp [h_antidistrib]
145+
146+
/-- Symmetric statement to `option.map₂_map_left_anticomm`. -/
147+
lemma map_map₂_antidistrib_left {g : γ → δ} {f' : β' → α → δ} {g' : β → β'}
148+
(h_antidistrib : ∀ a b, g (f a b) = f' (g' b) a) :
149+
(map₂ f a b).map g = map₂ f' (b.map g') a :=
150+
by cases a; cases b; simp [h_antidistrib]
151+
152+
/-- Symmetric statement to `option.map_map₂_right_anticomm`. -/
153+
lemma map_map₂_antidistrib_right {g : γ → δ} {f' : β → α' → δ} {g' : α → α'}
154+
(h_antidistrib : ∀ a b, g (f a b) = f' b (g' a)) :
155+
(map₂ f a b).map g = map₂ f' b (a.map g') :=
156+
by cases a; cases b; simp [h_antidistrib]
157+
158+
/-- Symmetric statement to `option.map_map₂_antidistrib_left`. -/
159+
lemma map₂_map_left_anticomm {f : α' → β → γ} {g : α → α'} {f' : β → α → δ} {g' : δ → γ}
160+
(h_left_anticomm : ∀ a b, f (g a) b = g' (f' b a)) :
161+
map₂ f (a.map g) b = (map₂ f' b a).map g' :=
162+
by cases a; cases b; simp [h_left_anticomm]
163+
164+
/-- Symmetric statement to `option.map_map₂_antidistrib_right`. -/
165+
lemma map_map₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : β → α → δ} {g' : δ → γ}
166+
(h_right_anticomm : ∀ a b, f a (g b) = g' (f' b a)) :
167+
map₂ f a (b.map g) = (map₂ f' b a).map g' :=
168+
by cases a; cases b; simp [h_right_anticomm]
169+
170+
end option

src/data/set/basic.lean

Lines changed: 10 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -2636,7 +2636,8 @@ end image_preimage
26362636
/-!
26372637
### Images of binary and ternary functions
26382638
2639-
This section is very similar to `order.filter.n_ary`. Please keep them in sync.
2639+
This section is very similar to `order.filter.n_ary`, `data.finset.n_ary`, `data.option.n_ary`.
2640+
Please keep them in sync.
26402641
-/
26412642

26422643
section n_ary_image
@@ -2834,25 +2835,25 @@ lemma image_image2_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ : α
28342835
(image2 f s t).image g = image2 f' (s.image g₁) (t.image g₂) :=
28352836
by simp_rw [image_image2, image2_image_left, image2_image_right, h_distrib]
28362837

2837-
/-- Symmetric of `set.image2_image_left_comm`. -/
2838+
/-- Symmetric statement to `set.image2_image_left_comm`. -/
28382839
lemma image_image2_distrib_left {g : γ → δ} {f' : α' → β → δ} {g' : α → α'}
28392840
(h_distrib : ∀ a b, g (f a b) = f' (g' a) b) :
28402841
(image2 f s t).image g = image2 f' (s.image g') t :=
28412842
(image_image2_distrib h_distrib).trans $ by rw image_id'
28422843

2843-
/-- Symmetric of `set.image_image2_right_comm`. -/
2844+
/-- Symmetric statement to `set.image_image2_right_comm`. -/
28442845
lemma image_image2_distrib_right {g : γ → δ} {f' : α → β' → δ} {g' : β → β'}
28452846
(h_distrib : ∀ a b, g (f a b) = f' a (g' b)) :
28462847
(image2 f s t).image g = image2 f' s (t.image g') :=
28472848
(image_image2_distrib h_distrib).trans $ by rw image_id'
28482849

2849-
/-- Symmetric of `set.image_image2_distrib_left`. -/
2850+
/-- Symmetric statement to `set.image_image2_distrib_left`. -/
28502851
lemma image2_image_left_comm {f : α' → β → γ} {g : α → α'} {f' : α → β → δ} {g' : δ → γ}
28512852
(h_left_comm : ∀ a b, f (g a) b = g' (f' a b)) :
28522853
image2 f (s.image g) t = (image2 f' s t).image g' :=
28532854
(image_image2_distrib_left $ λ a b, (h_left_comm a b).symm).symm
28542855

2855-
/-- Symmetric of `set.image_image2_distrib_right`. -/
2856+
/-- Symmetric statement to `set.image_image2_distrib_right`. -/
28562857
lemma image_image2_right_comm {f : α → β' → γ} {g : β → β'} {f' : α → β → δ} {g' : δ → γ}
28572858
(h_right_comm : ∀ a b, f a (g b) = g' (f' a b)) :
28582859
image2 f s (t.image g) = (image2 f' s t).image g' :=
@@ -2883,25 +2884,25 @@ lemma image_image2_antidistrib {g : γ → δ} {f' : β' → α' → δ} {g₁ :
28832884
(image2 f s t).image g = image2 f' (t.image g₁) (s.image g₂) :=
28842885
by { rw image2_swap f, exact image_image2_distrib (λ _ _, h_antidistrib _ _) }
28852886

2886-
/-- Symmetric of `set.image2_image_left_anticomm`. -/
2887+
/-- Symmetric statement to `set.image2_image_left_anticomm`. -/
28872888
lemma image_image2_antidistrib_left {g : γ → δ} {f' : β' → α → δ} {g' : β → β'}
28882889
(h_antidistrib : ∀ a b, g (f a b) = f' (g' b) a) :
28892890
(image2 f s t).image g = image2 f' (t.image g') s :=
28902891
(image_image2_antidistrib h_antidistrib).trans $ by rw image_id'
28912892

2892-
/-- Symmetric of `set.image_image2_right_anticomm`. -/
2893+
/-- Symmetric statement to `set.image_image2_right_anticomm`. -/
28932894
lemma image_image2_antidistrib_right {g : γ → δ} {f' : β → α' → δ} {g' : α → α'}
28942895
(h_antidistrib : ∀ a b, g (f a b) = f' b (g' a)) :
28952896
(image2 f s t).image g = image2 f' t (s.image g') :=
28962897
(image_image2_antidistrib h_antidistrib).trans $ by rw image_id'
28972898

2898-
/-- Symmetric of `set.image_image2_antidistrib_left`. -/
2899+
/-- Symmetric statement to `set.image_image2_antidistrib_left`. -/
28992900
lemma image2_image_left_anticomm {f : α' → β → γ} {g : α → α'} {f' : β → α → δ} {g' : δ → γ}
29002901
(h_left_anticomm : ∀ a b, f (g a) b = g' (f' b a)) :
29012902
image2 f (s.image g) t = (image2 f' t s).image g' :=
29022903
(image_image2_antidistrib_left $ λ a b, (h_left_anticomm b a).symm).symm
29032904

2904-
/-- Symmetric of `set.image_image2_antidistrib_right`. -/
2905+
/-- Symmetric statement to `set.image_image2_antidistrib_right`. -/
29052906
lemma image_image2_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : β → α → δ} {g' : δ → γ}
29062907
(h_right_anticomm : ∀ a b, f a (g b) = g' (f' b a)) :
29072908
image2 f s (t.image g) = (image2 f' t s).image g' :=

src/data/set/functor.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -29,6 +29,12 @@ instance : monad.{u} set :=
2929
@[simp] lemma seq_eq_set_seq (s : set (α → β)) (t : set α) : s <*> t = s.seq t := rfl
3030
@[simp] lemma pure_def (a : α) : (pure a : set α) = {a} := rfl
3131

32+
/-- `set.image2` in terms of monadic operations. Note that this can't be taken as the definition
33+
because of the lack of universe polymorphism. -/
34+
lemma image2_def {α β γ : Type*} (f : α → β → γ) (s : set α) (t : set β) :
35+
image2 f s t = f <$> s <*> t :=
36+
by { ext, simp }
37+
3238
instance : is_lawful_monad set :=
3339
{ id_map := λ α, image_id,
3440
comp_map := λ α β γ f g s, image_comp _ _ _,

src/data/set/prod.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -281,6 +281,13 @@ set.ext $ λ a,
281281

282282
@[simp] lemma image2_mk_eq_prod : image2 prod.mk s t = s ×ˢ t := ext $ by simp
283283

284+
@[simp] lemma image2_curry (f : α × β → γ) (s : set α) (t : set β) :
285+
image2 (λ a b, f (a, b)) s t = (s ×ˢ t).image f :=
286+
by rw [←image2_mk_eq_prod, image_image2]
287+
288+
@[simp] lemma image_uncurry_prod (f : α → β → γ) (s : set α) (t : set β) :
289+
uncurry f '' s ×ˢ t = image2 f s t := by { rw ←image2_curry, refl }
290+
284291
section mono
285292

286293
variables [preorder α] {f : α → set β} {g : α → set γ}

0 commit comments

Comments
 (0)