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,
⋃ i, s i
is callediUnion
⋂ i, s i
is callediInter
⋃ i j, s i j
is callediUnion₂
. This is aniUnion
inside aniUnion
.⋂ i j, s i j
is callediInter₂
. This is aniInter
inside aniInter
.⋃ i ∈ s, t i
is calledbiUnion
for "boundediUnion
". This is the special case ofiUnion₂
wherej : i ∈ s
.⋂ i ∈ s, t i
is calledbiInter
for "boundediInter
". This is the special case ofiInter₂
wherej : i ∈ s
.
Notation #
⋃
:Set.iUnion
⋂
:Set.iInter
⋃₀
:Set.sUnion
⋂₀
:Set.sInter
theorem
Set.image_preimage
{α : Type u_1}
{β : Type u_2}
{f : α → β}
:
GaloisConnection (image f) (preimage f)
theorem
Set.preimage_kernImage
{α : Type u_1}
{β : Type u_2}
{f : α → β}
:
GaloisConnection (preimage f) (kernImage f)
Bounded unions and intersections #
Lemmas about Set.MapsTo
#
restrictPreimage
#
InjOn
#
theorem
Set.image_iInter
{α : Type u_1}
{β : Type u_2}
{ι : Sort u_5}
{f : α → β}
(hf : Function.Bijective f)
(s : ι → Set α)
:
SurjOn
#
BijOn
#
image
, preimage
#
theorem
Set.image2_eq_iUnion
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
(f : α → β → γ)
(s : Set α)
(t : Set β)
:
The Set.image2
version of Set.image_eq_iUnion