Functions over sets #
This file contains basic results on the following predicates of functions and sets:
Set.EqOn f₁ f₂ s
: functionsf₁
andf₂
are equal at every point ofs
;Set.MapsTo f s t
:f
sends every point ofs
to a point oft
;Set.InjOn f s
: restriction off
tos
is injective;Set.SurjOn f s t
: every point ins
has a preimage ins
;Set.BijOn f s t
:f
is a bijection betweens
andt
;Set.LeftInvOn f' f s
: for everyx ∈ s
we havef' (f x) = x
;Set.RightInvOn f' f t
: for everyy ∈ t
we havef (f' y) = y
;Set.InvOn f' f s t
:f'
is a two-side inverse off
ons
andt
, i.e. we haveSet.LeftInvOn f' f s
andSet.RightInvOn f' f t
.
Equality on a set #
Variant of EqOn.image_eq
, for one function being the identity.
Injectivity on a set #
Alias of Set.injOn_of_injective
.
Surjectivity on a set #
Bijectivity #
If we have a commutative square
α --f--> β
| |
p₁ p₂
| |
\/ \/
γ --g--> δ
and f
induces a bijection from s : Set α
to t : Set β
, then g
induces a bijection from the image of s
to the image of t
, as long as g
is
is injective on the image of s
.
Alias of the forward direction of Set.bijective_iff_bijOn_univ
.
left inverse #
Right inverse #
Two-side inverses #
If functions f'
and f
are inverse on s
and t
, f
maps s
into t
, and f'
maps t
into s
, then f
is a bijection between s
and t
. The mapsTo
arguments can be deduced from
surjOn
statements using LeftInvOn.mapsTo
and RightInvOn.mapsTo
.
Construct the inverse for a function f
on domain s
. This function is a right inverse of f
on f '' s
. For a computable version, see Function.Embedding.invOfMemRange
.
Equations
- Function.invFunOn f s b = if h : ∃ a ∈ s, f a = b then Classical.choose h else Classical.choice inst✝
Instances For
This lemma is a special case of rightInvOn_invFunOn.image_image'
; it may make more sense
to use the other lemma directly in an application.
This lemma is a special case of rightInvOn_invFunOn.image_image
; it may make more sense
to use the other lemma directly in an application.
Alias of the forward direction of Set.surjOn_iff_exists_bijOn_subset
.
If f
maps s
bijectively to t
and a set t'
is contained in the image of some s₁ ⊇ s
,
then s₁
has a subset containing s
that f
maps bijectively to t'
.
If f
maps s
bijectively to t
, and t'
is a superset of t
contained in the range of f
,
then f
maps some superset of s
bijectively to t'
.
Non-dependent version of Function.update_comp_eq_of_not_mem_range'