Documentation

Mathlib.Data.Set.Lattice.Image

The set lattice and (pre)images of functions #

This file contains lemmas on the interaction between the indexed union/intersection of sets and the image and preimage operations: Set.image, Set.preimage, Set.image2, Set.kernImage. It also covers Set.MapsTo, Set.InjOn, Set.SurjOn, Set.BijOn.

In order to accommodate Set.image2, the file includes results on union/intersection in products.

Naming convention #

In lemma names,

Notation #

theorem Set.image_preimage {α : Type u_1} {β : Type u_2} {f : αβ} :
theorem Set.preimage_kernImage {α : Type u_1} {β : Type u_2} {f : αβ} :
theorem Set.kernImage_mono {α : Type u_1} {β : Type u_2} {f : αβ} :
theorem Set.kernImage_eq_compl {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set α} :
kernImage f s = (f '' s)
theorem Set.kernImage_compl {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set α} :
kernImage f s = (f '' s)
theorem Set.kernImage_empty {α : Type u_1} {β : Type u_2} {f : αβ} :
theorem Set.kernImage_preimage_eq_iff {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set β} :
kernImage f (f ⁻¹' s) = s (range f) s
theorem Set.compl_range_subset_kernImage {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set α} :
theorem Set.kernImage_union_preimage {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set α} {t : Set β} :
kernImage f (s f ⁻¹' t) = kernImage f s t
theorem Set.kernImage_preimage_union {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set α} {t : Set β} :
kernImage f (f ⁻¹' t s) = t kernImage f s
theorem Set.image_projection_prod {ι : Type u_9} {α : ιType u_10} {v : (i : ι) → Set (α i)} (hv : (univ.pi v).Nonempty) (i : ι) :
(fun (x : (i : ι) → α i) => x i) '' ⋂ (k : ι), (fun (x : (j : ι) → α j) => x k) ⁻¹' v k = v i

Bounded unions and intersections #

Lemmas about Set.MapsTo #

@[simp]
theorem Set.mapsTo_sUnion {α : Type u_1} {β : Type u_2} {S : Set (Set α)} {t : Set β} {f : αβ} :
MapsTo f (⋃₀ S) t sS, MapsTo f s t
@[simp]
theorem Set.mapsTo_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : Set β} {f : αβ} :
MapsTo f (⋃ (i : ι), s i) t ∀ (i : ι), MapsTo f (s i) t
theorem Set.mapsTo_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : (i : ι) → κ iSet α} {t : Set β} {f : αβ} :
MapsTo f (⋃ (i : ι), ⋃ (j : κ i), s i j) t ∀ (i : ι) (j : κ i), MapsTo f (s i j) t
theorem Set.mapsTo_iUnion_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), MapsTo f (s i) (t i)) :
MapsTo f (⋃ (i : ι), s i) (⋃ (i : ι), t i)
theorem Set.mapsTo_iUnion₂_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : (i : ι) → κ iSet α} {t : (i : ι) → κ iSet β} {f : αβ} (H : ∀ (i : ι) (j : κ i), MapsTo f (s i j) (t i j)) :
MapsTo f (⋃ (i : ι), ⋃ (j : κ i), s i j) (⋃ (i : ι), ⋃ (j : κ i), t i j)
@[simp]
theorem Set.mapsTo_sInter {α : Type u_1} {β : Type u_2} {s : Set α} {T : Set (Set β)} {f : αβ} :
MapsTo f s (⋂₀ T) tT, MapsTo f s t
@[simp]
theorem Set.mapsTo_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : Set α} {t : ιSet β} {f : αβ} :
MapsTo f s (⋂ (i : ι), t i) ∀ (i : ι), MapsTo f s (t i)
theorem Set.mapsTo_iInter₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : Set α} {t : (i : ι) → κ iSet β} {f : αβ} :
MapsTo f s (⋂ (i : ι), ⋂ (j : κ i), t i j) ∀ (i : ι) (j : κ i), MapsTo f s (t i j)
theorem Set.mapsTo_iInter_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), MapsTo f (s i) (t i)) :
MapsTo f (⋂ (i : ι), s i) (⋂ (i : ι), t i)
theorem Set.mapsTo_iInter₂_iInter₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : (i : ι) → κ iSet α} {t : (i : ι) → κ iSet β} {f : αβ} (H : ∀ (i : ι) (j : κ i), MapsTo f (s i j) (t i j)) :
MapsTo f (⋂ (i : ι), ⋂ (j : κ i), s i j) (⋂ (i : ι), ⋂ (j : κ i), t i j)
theorem Set.image_iInter_subset {α : Type u_1} {β : Type u_2} {ι : Sort u_5} (s : ιSet α) (f : αβ) :
f '' ⋂ (i : ι), s i ⋂ (i : ι), f '' s i
theorem Set.image_iInter₂_subset {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} (s : (i : ι) → κ iSet α) (f : αβ) :
f '' ⋂ (i : ι), ⋂ (j : κ i), s i j ⋂ (i : ι), ⋂ (j : κ i), f '' s i j
theorem Set.image_sInter_subset {α : Type u_1} {β : Type u_2} (S : Set (Set α)) (f : αβ) :
f '' ⋂₀ S sS, f '' s
theorem Set.image2_sInter_right_subset {α : Type u_1} {β : Type u_2} {γ : Type u_3} (t : Set α) (S : Set (Set β)) (f : αβγ) :
image2 f t (⋂₀ S) sS, image2 f t s
theorem Set.image2_sInter_left_subset {α : Type u_1} {β : Type u_2} {γ : Type u_3} (S : Set (Set α)) (t : Set β) (f : αβγ) :
image2 f (⋂₀ S) t sS, image2 f s t

restrictPreimage #

theorem Set.injective_iff_injective_of_iUnion_eq_univ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {U : ιSet β} (hU : iUnion U = univ) :
theorem Set.surjective_iff_surjective_of_iUnion_eq_univ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {U : ιSet β} (hU : iUnion U = univ) :
theorem Set.bijective_iff_bijective_of_iUnion_eq_univ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {U : ιSet β} (hU : iUnion U = univ) :

InjOn #

theorem Set.InjOn.image_iInter_eq {α : Type u_1} {β : Type u_2} {ι : Sort u_5} [Nonempty ι] {s : ιSet α} {f : αβ} (h : InjOn f (⋃ (i : ι), s i)) :
f '' ⋂ (i : ι), s i = ⋂ (i : ι), f '' s i
theorem Set.InjOn.image_biInter_eq {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {p : ιProp} {s : (i : ι) → p iSet α} (hp : ∃ (i : ι), p i) {f : αβ} (h : InjOn f (⋃ (i : ι), ⋃ (hi : p i), s i hi)) :
f '' ⋂ (i : ι), ⋂ (hi : p i), s i hi = ⋂ (i : ι), ⋂ (hi : p i), f '' s i hi
theorem Set.image_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} (hf : Function.Bijective f) (s : ιSet α) :
f '' ⋂ (i : ι), s i = ⋂ (i : ι), f '' s i
theorem Set.image_iInter₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {f : αβ} (hf : Function.Bijective f) (s : (i : ι) → κ iSet α) :
f '' ⋂ (i : ι), ⋂ (j : κ i), s i j = ⋂ (i : ι), ⋂ (j : κ i), f '' s i j
theorem Set.inj_on_iUnion_of_directed {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} (hs : Directed (fun (x1 x2 : Set α) => x1 x2) s) {f : αβ} (hf : ∀ (i : ι), InjOn f (s i)) :
InjOn f (⋃ (i : ι), s i)

SurjOn #

theorem Set.surjOn_sUnion {α : Type u_1} {β : Type u_2} {s : Set α} {T : Set (Set β)} {f : αβ} (H : tT, SurjOn f s t) :
SurjOn f s (⋃₀ T)
theorem Set.surjOn_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : Set α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), SurjOn f s (t i)) :
SurjOn f s (⋃ (i : ι), t i)
theorem Set.surjOn_iUnion_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), SurjOn f (s i) (t i)) :
SurjOn f (⋃ (i : ι), s i) (⋃ (i : ι), t i)
theorem Set.surjOn_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : Set α} {t : (i : ι) → κ iSet β} {f : αβ} (H : ∀ (i : ι) (j : κ i), SurjOn f s (t i j)) :
SurjOn f s (⋃ (i : ι), ⋃ (j : κ i), t i j)
theorem Set.surjOn_iUnion₂_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : (i : ι) → κ iSet α} {t : (i : ι) → κ iSet β} {f : αβ} (H : ∀ (i : ι) (j : κ i), SurjOn f (s i j) (t i j)) :
SurjOn f (⋃ (i : ι), ⋃ (j : κ i), s i j) (⋃ (i : ι), ⋃ (j : κ i), t i j)
theorem Set.surjOn_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} [Nonempty ι] {s : ιSet α} {t : Set β} {f : αβ} (H : ∀ (i : ι), SurjOn f (s i) t) (Hinj : InjOn f (⋃ (i : ι), s i)) :
SurjOn f (⋂ (i : ι), s i) t
theorem Set.surjOn_iInter_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} [Nonempty ι] {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), SurjOn f (s i) (t i)) (Hinj : InjOn f (⋃ (i : ι), s i)) :
SurjOn f (⋂ (i : ι), s i) (⋂ (i : ι), t i)

BijOn #

theorem Set.bijOn_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), BijOn f (s i) (t i)) (Hinj : InjOn f (⋃ (i : ι), s i)) :
BijOn f (⋃ (i : ι), s i) (⋃ (i : ι), t i)
theorem Set.bijOn_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} [hi : Nonempty ι] {s : ιSet α} {t : ιSet β} {f : αβ} (H : ∀ (i : ι), BijOn f (s i) (t i)) (Hinj : InjOn f (⋃ (i : ι), s i)) :
BijOn f (⋂ (i : ι), s i) (⋂ (i : ι), t i)
theorem Set.bijOn_iUnion_of_directed {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} (hs : Directed (fun (x1 x2 : Set α) => x1 x2) s) {t : ιSet β} {f : αβ} (H : ∀ (i : ι), BijOn f (s i) (t i)) :
BijOn f (⋃ (i : ι), s i) (⋃ (i : ι), t i)
theorem Set.bijOn_iInter_of_directed {α : Type u_1} {β : Type u_2} {ι : Sort u_5} [Nonempty ι] {s : ιSet α} (hs : Directed (fun (x1 x2 : Set α) => x1 x2) s) {t : ιSet β} {f : αβ} (H : ∀ (i : ι), BijOn f (s i) (t i)) :
BijOn f (⋂ (i : ι), s i) (⋂ (i : ι), t i)

image, preimage #

theorem Set.image_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {s : ιSet α} :
f '' ⋃ (i : ι), s i = ⋃ (i : ι), f '' s i
theorem Set.image_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} (f : αβ) (s : (i : ι) → κ iSet α) :
f '' ⋃ (i : ι), ⋃ (j : κ i), s i j = ⋃ (i : ι), ⋃ (j : κ i), f '' s i j
theorem Set.univ_subtype {α : Type u_1} {p : αProp} :
univ = ⋃ (x : α), ⋃ (h : p x), {x, h}
theorem Set.range_eq_iUnion {α : Type u_1} {ι : Sort u_9} (f : ια) :
range f = ⋃ (i : ι), {f i}
theorem Set.image_eq_iUnion {α : Type u_1} {β : Type u_2} (f : αβ) (s : Set α) :
f '' s = is, {f i}
theorem Set.biUnion_range {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : ια} {g : αSet β} :
xrange f, g x = ⋃ (y : ι), g (f y)
@[simp]
theorem Set.iUnion_iUnion_eq' {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : ια} {g : αSet β} :
⋃ (x : α), ⋃ (y : ι), ⋃ (_ : f y = x), g x = ⋃ (y : ι), g (f y)
theorem Set.biInter_range {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : ια} {g : αSet β} :
xrange f, g x = ⋂ (y : ι), g (f y)
@[simp]
theorem Set.iInter_iInter_eq' {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : ια} {g : αSet β} :
⋂ (x : α), ⋂ (y : ι), ⋂ (_ : f y = x), g x = ⋂ (y : ι), g (f y)
theorem Set.biUnion_image {α : Type u_1} {β : Type u_2} {γ : Type u_3} {s : Set γ} {f : γα} {g : αSet β} :
xf '' s, g x = ys, g (f y)
theorem Set.biInter_image {α : Type u_1} {β : Type u_2} {γ : Type u_3} {s : Set γ} {f : γα} {g : αSet β} :
xf '' s, g x = ys, g (f y)
theorem Set.biUnion_image2 {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (s : Set α) (t : Set β) (f : αβγ) (g : γSet δ) :
cimage2 f s t, g c = as, bt, g (f a b)
theorem Set.biInter_image2 {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} (s : Set α) (t : Set β) (f : αβγ) (g : γSet δ) :
cimage2 f s t, g c = as, bt, g (f a b)
theorem Set.iUnion_inter_iUnion {α : Type u_1} {ι : Sort u_9} {κ : Sort u_10} (f : ιSet α) (g : κSet α) :
(⋃ (i : ι), f i) ⋃ (j : κ), g j = ⋃ (i : ι), ⋃ (j : κ), f i g j
theorem Set.iInter_union_iInter {α : Type u_1} {ι : Sort u_9} {κ : Sort u_10} (f : ιSet α) (g : κSet α) :
(⋂ (i : ι), f i) ⋂ (j : κ), g j = ⋂ (i : ι), ⋂ (j : κ), f i g j
theorem Set.iUnion₂_inter_iUnion₂ {α : Type u_1} {ι₁ : Sort u_9} {κ₁ : Sort u_10} {ι₂ : ι₁Sort u_11} {k₂ : κ₁Sort u_12} (f : (i₁ : ι₁) → ι₂ i₁Set α) (g : (j₁ : κ₁) → k₂ j₁Set α) :
(⋃ (i₁ : ι₁), ⋃ (i₂ : ι₂ i₁), f i₁ i₂) ⋃ (j₁ : κ₁), ⋃ (j₂ : k₂ j₁), g j₁ j₂ = ⋃ (i₁ : ι₁), ⋃ (i₂ : ι₂ i₁), ⋃ (j₁ : κ₁), ⋃ (j₂ : k₂ j₁), f i₁ i₂ g j₁ j₂
theorem Set.iInter₂_union_iInter₂ {α : Type u_1} {ι₁ : Sort u_9} {κ₁ : Sort u_10} {ι₂ : ι₁Sort u_11} {k₂ : κ₁Sort u_12} (f : (i₁ : ι₁) → ι₂ i₁Set α) (g : (j₁ : κ₁) → k₂ j₁Set α) :
(⋂ (i₁ : ι₁), ⋂ (i₂ : ι₂ i₁), f i₁ i₂) ⋂ (j₁ : κ₁), ⋂ (j₂ : k₂ j₁), g j₁ j₂ = ⋂ (i₁ : ι₁), ⋂ (i₂ : ι₂ i₁), ⋂ (j₁ : κ₁), ⋂ (j₂ : k₂ j₁), f i₁ i₂ g j₁ j₂
theorem Set.monotone_preimage {α : Type u_1} {β : Type u_2} {f : αβ} :
@[simp]
theorem Set.preimage_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {s : ιSet β} :
f ⁻¹' ⋃ (i : ι), s i = ⋃ (i : ι), f ⁻¹' s i
theorem Set.preimage_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {f : αβ} {s : (i : ι) → κ iSet β} :
f ⁻¹' ⋃ (i : ι), ⋃ (j : κ i), s i j = ⋃ (i : ι), ⋃ (j : κ i), f ⁻¹' s i j
theorem Set.image_sUnion {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set (Set α)} :
f '' ⋃₀ s = ⋃₀ (image f '' s)
@[simp]
theorem Set.preimage_sUnion {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set (Set β)} :
f ⁻¹' ⋃₀ s = ts, f ⁻¹' t
theorem Set.preimage_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {f : αβ} {s : ιSet β} :
f ⁻¹' ⋂ (i : ι), s i = ⋂ (i : ι), f ⁻¹' s i
theorem Set.preimage_iInter₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {f : αβ} {s : (i : ι) → κ iSet β} :
f ⁻¹' ⋂ (i : ι), ⋂ (j : κ i), s i j = ⋂ (i : ι), ⋂ (j : κ i), f ⁻¹' s i j
@[simp]
theorem Set.preimage_sInter {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set (Set β)} :
f ⁻¹' ⋂₀ s = ts, f ⁻¹' t
@[simp]
theorem Set.biUnion_preimage_singleton {α : Type u_1} {β : Type u_2} (f : αβ) (s : Set β) :
ys, f ⁻¹' {y} = f ⁻¹' s
theorem Set.biUnion_range_preimage_singleton {α : Type u_1} {β : Type u_2} (f : αβ) :
yrange f, f ⁻¹' {y} = univ
theorem Set.prod_iUnion {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : Set α} {t : ιSet β} :
s ×ˢ ⋃ (i : ι), t i = ⋃ (i : ι), s ×ˢ t i
theorem Set.prod_iUnion₂ {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : Set α} {t : (i : ι) → κ iSet β} :
s ×ˢ ⋃ (i : ι), ⋃ (j : κ i), t i j = ⋃ (i : ι), ⋃ (j : κ i), s ×ˢ t i j
theorem Set.prod_sUnion {α : Type u_1} {β : Type u_2} {s : Set α} {C : Set (Set β)} :
s ×ˢ ⋃₀ C = ⋃₀ ((fun (t : Set β) => s ×ˢ t) '' C)
theorem Set.iUnion_prod_const {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : ιSet α} {t : Set β} :
(⋃ (i : ι), s i) ×ˢ t = ⋃ (i : ι), s i ×ˢ t
theorem Set.iUnion₂_prod_const {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {κ : ιSort u_8} {s : (i : ι) → κ iSet α} {t : Set β} :
(⋃ (i : ι), ⋃ (j : κ i), s i j) ×ˢ t = ⋃ (i : ι), ⋃ (j : κ i), s i j ×ˢ t
theorem Set.sUnion_prod_const {α : Type u_1} {β : Type u_2} {C : Set (Set α)} {t : Set β} :
⋃₀ C ×ˢ t = ⋃₀ ((fun (s : Set α) => s ×ˢ t) '' C)
theorem Set.iUnion_prod {ι : Type u_9} {ι' : Type u_10} {α : Type u_11} {β : Type u_12} (s : ιSet α) (t : ι'Set β) :
⋃ (x : ι × ι'), s x.1 ×ˢ t x.2 = (⋃ (i : ι), s i) ×ˢ ⋃ (i : ι'), t i
theorem Set.iUnion_prod' {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : β × γSet α) :
⋃ (x : β × γ), f x = ⋃ (i : β), ⋃ (j : γ), f (i, j)

Analogue of iSup_prod for sets.

theorem Set.iUnion_prod_of_monotone {α : Type u_1} {β : Type u_2} {γ : Type u_3} [SemilatticeSup α] {s : αSet β} {t : αSet γ} (hs : Monotone s) (ht : Monotone t) :
⋃ (x : α), s x ×ˢ t x = (⋃ (x : α), s x) ×ˢ ⋃ (x : α), t x
theorem Set.sInter_prod_sInter_subset {α : Type u_1} {β : Type u_2} (S : Set (Set α)) (T : Set (Set β)) :
⋂₀ S ×ˢ ⋂₀ T rS ×ˢ T, r.1 ×ˢ r.2
theorem Set.sInter_prod_sInter {α : Type u_1} {β : Type u_2} {S : Set (Set α)} {T : Set (Set β)} (hS : S.Nonempty) (hT : T.Nonempty) :
⋂₀ S ×ˢ ⋂₀ T = rS ×ˢ T, r.1 ×ˢ r.2
theorem Set.sInter_prod {α : Type u_1} {β : Type u_2} {S : Set (Set α)} (hS : S.Nonempty) (t : Set β) :
⋂₀ S ×ˢ t = sS, s ×ˢ t
theorem Set.prod_sInter {α : Type u_1} {β : Type u_2} {T : Set (Set β)} (hT : T.Nonempty) (s : Set α) :
s ×ˢ ⋂₀ T = tT, s ×ˢ t
theorem Set.prod_iInter {α : Type u_1} {β : Type u_2} {ι : Sort u_5} {s : Set α} {t : ιSet β} [ : Nonempty ι] :
s ×ˢ ⋂ (i : ι), t i = ⋂ (i : ι), s ×ˢ t i
theorem Set.image2_eq_iUnion {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (s : Set α) (t : Set β) :
image2 f s t = is, jt, {f i j}

The Set.image2 version of Set.image_eq_iUnion

theorem Set.iUnion_image_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) {s : Set α} {t : Set β} :
as, f a '' t = image2 f s t
theorem Set.iUnion_image_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) {s : Set α} {t : Set β} :
bt, (fun (x : α) => f x b) '' s = image2 f s t
theorem Set.image2_iUnion_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} (f : αβγ) (s : ιSet α) (t : Set β) :
image2 f (⋃ (i : ι), s i) t = ⋃ (i : ι), image2 f (s i) t
theorem Set.image2_iUnion_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} (f : αβγ) (s : Set α) (t : ιSet β) :
image2 f s (⋃ (i : ι), t i) = ⋃ (i : ι), image2 f s (t i)
theorem Set.image2_sUnion_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (S : Set (Set α)) (t : Set β) :
image2 f (⋃₀ S) t = sS, image2 f s t
theorem Set.image2_sUnion_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (s : Set α) (T : Set (Set β)) :
image2 f s (⋃₀ T) = tT, image2 f s t
theorem Set.image2_iUnion₂_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} {κ : ιSort u_8} (f : αβγ) (s : (i : ι) → κ iSet α) (t : Set β) :
image2 f (⋃ (i : ι), ⋃ (j : κ i), s i j) t = ⋃ (i : ι), ⋃ (j : κ i), image2 f (s i j) t
theorem Set.image2_iUnion₂_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} {κ : ιSort u_8} (f : αβγ) (s : Set α) (t : (i : ι) → κ iSet β) :
image2 f s (⋃ (i : ι), ⋃ (j : κ i), t i j) = ⋃ (i : ι), ⋃ (j : κ i), image2 f s (t i j)
theorem Set.image2_iInter_subset_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} (f : αβγ) (s : ιSet α) (t : Set β) :
image2 f (⋂ (i : ι), s i) t ⋂ (i : ι), image2 f (s i) t
theorem Set.image2_iInter_subset_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} (f : αβγ) (s : Set α) (t : ιSet β) :
image2 f s (⋂ (i : ι), t i) ⋂ (i : ι), image2 f s (t i)
theorem Set.image2_iInter₂_subset_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} {κ : ιSort u_8} (f : αβγ) (s : (i : ι) → κ iSet α) (t : Set β) :
image2 f (⋂ (i : ι), ⋂ (j : κ i), s i j) t ⋂ (i : ι), ⋂ (j : κ i), image2 f (s i j) t
theorem Set.image2_iInter₂_subset_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Sort u_5} {κ : ιSort u_8} (f : αβγ) (s : Set α) (t : (i : ι) → κ iSet β) :
image2 f s (⋂ (i : ι), ⋂ (j : κ i), t i j) ⋂ (i : ι), ⋂ (j : κ i), image2 f s (t i j)
theorem Set.image2_sInter_subset_left {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (S : Set (Set α)) (t : Set β) :
image2 f (⋂₀ S) t sS, image2 f s t
theorem Set.image2_sInter_subset_right {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (s : Set α) (T : Set (Set β)) :
image2 f s (⋂₀ T) tT, image2 f s t
theorem Set.prod_eq_biUnion_left {α : Type u_1} {β : Type u_2} {s : Set α} {t : Set β} :
s ×ˢ t = as, (fun (b : β) => (a, b)) '' t
theorem Set.prod_eq_biUnion_right {α : Type u_1} {β : Type u_2} {s : Set α} {t : Set β} :
s ×ˢ t = bt, (fun (a : α) => (a, b)) '' s
theorem Set.seq_def {α : Type u_1} {β : Type u_2} {s : Set (αβ)} {t : Set α} :
s.seq t = fs, f '' t
theorem Set.seq_subset {α : Type u_1} {β : Type u_2} {s : Set (αβ)} {t : Set α} {u : Set β} :
s.seq t u fs, at, f a u
theorem Set.seq_mono {α : Type u_1} {β : Type u_2} {s₀ s₁ : Set (αβ)} {t₀ t₁ : Set α} (hs : s₀ s₁) (ht : t₀ t₁) :
s₀.seq t₀ s₁.seq t₁
theorem Set.singleton_seq {α : Type u_1} {β : Type u_2} {f : αβ} {t : Set α} :
{f}.seq t = f '' t
theorem Set.seq_singleton {α : Type u_1} {β : Type u_2} {s : Set (αβ)} {a : α} :
s.seq {a} = (fun (f : αβ) => f a) '' s
theorem Set.seq_seq {α : Type u_1} {β : Type u_2} {γ : Type u_3} {s : Set (βγ)} {t : Set (αβ)} {u : Set α} :
s.seq (t.seq u) = (((fun (x1 : βγ) (x2 : αβ) => x1 x2) '' s).seq t).seq u
theorem Set.image_seq {α : Type u_1} {β : Type u_2} {γ : Type u_3} {f : βγ} {s : Set (αβ)} {t : Set α} :
f '' s.seq t = ((fun (x : αβ) => f x) '' s).seq t
theorem Set.prod_eq_seq {α : Type u_1} {β : Type u_2} {s : Set α} {t : Set β} :
s ×ˢ t = (Prod.mk '' s).seq t
theorem Set.prod_image_seq_comm {α : Type u_1} {β : Type u_2} (s : Set α) (t : Set β) :
(Prod.mk '' s).seq t = ((fun (b : β) (a : α) => (a, b)) '' t).seq s
theorem Set.image2_eq_seq {α : Type u_1} {β : Type u_2} {γ : Type u_3} (f : αβγ) (s : Set α) (t : Set β) :
image2 f s t = (f '' s).seq t