Type | Label | Description |
Statement |
|
Theorem | funimaex 5301 |
The image of a set under any function is also a set. Equivalent of
Axiom of Replacement. Axiom 39(vi) of [Quine] p. 284. Compare Exercise
9 of [TakeutiZaring] p. 29.
(Contributed by NM, 17-Nov-2002.)
|
β’ π΅ β V β β’ (Fun π΄ β (π΄ β π΅) β V) |
|
Theorem | isarep1 5302* |
Part of a study of the Axiom of Replacement used by the Isabelle prover.
The object PrimReplace is apparently the image of the function encoded
by π(π₯, π¦) i.e. the class ({β¨π₯, π¦β© β£ π} β π΄).
If so, we can prove Isabelle's "Axiom of Replacement"
conclusion without
using the Axiom of Replacement, for which I (N. Megill) currently have
no explanation. (Contributed by NM, 26-Oct-2006.) (Proof shortened by
Mario Carneiro, 4-Dec-2016.)
|
β’ (π β ({β¨π₯, π¦β© β£ π} β π΄) β βπ₯ β π΄ [π / π¦]π) |
|
Theorem | isarep2 5303* |
Part of a study of the Axiom of Replacement used by the Isabelle prover.
In Isabelle, the sethood of PrimReplace is apparently postulated
implicitly by its type signature "[ i,
[ i, i ] => o ]
=> i", which automatically asserts that it is a set without
using any
axioms. To prove that it is a set in Metamath, we need the hypotheses
of Isabelle's "Axiom of Replacement" as well as the Axiom of
Replacement
in the form funimaex 5301. (Contributed by NM, 26-Oct-2006.)
|
β’ π΄ β V & β’ βπ₯ β π΄ βπ¦βπ§((π β§ [π§ / π¦]π) β π¦ = π§) β β’ βπ€ π€ = ({β¨π₯, π¦β© β£ π} β π΄) |
|
Theorem | fneq1 5304 |
Equality theorem for function predicate with domain. (Contributed by NM,
1-Aug-1994.)
|
β’ (πΉ = πΊ β (πΉ Fn π΄ β πΊ Fn π΄)) |
|
Theorem | fneq2 5305 |
Equality theorem for function predicate with domain. (Contributed by NM,
1-Aug-1994.)
|
β’ (π΄ = π΅ β (πΉ Fn π΄ β πΉ Fn π΅)) |
|
Theorem | fneq1d 5306 |
Equality deduction for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
β’ (π β πΉ = πΊ) β β’ (π β (πΉ Fn π΄ β πΊ Fn π΄)) |
|
Theorem | fneq2d 5307 |
Equality deduction for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
β’ (π β π΄ = π΅) β β’ (π β (πΉ Fn π΄ β πΉ Fn π΅)) |
|
Theorem | fneq12d 5308 |
Equality deduction for function predicate with domain. (Contributed by
NM, 26-Jun-2011.)
|
β’ (π β πΉ = πΊ)
& β’ (π β π΄ = π΅) β β’ (π β (πΉ Fn π΄ β πΊ Fn π΅)) |
|
Theorem | fneq12 5309 |
Equality theorem for function predicate with domain. (Contributed by
Thierry Arnoux, 31-Jan-2017.)
|
β’ ((πΉ = πΊ β§ π΄ = π΅) β (πΉ Fn π΄ β πΊ Fn π΅)) |
|
Theorem | fneq1i 5310 |
Equality inference for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
β’ πΉ = πΊ β β’ (πΉ Fn π΄ β πΊ Fn π΄) |
|
Theorem | fneq2i 5311 |
Equality inference for function predicate with domain. (Contributed by
NM, 4-Sep-2011.)
|
β’ π΄ = π΅ β β’ (πΉ Fn π΄ β πΉ Fn π΅) |
|
Theorem | nffn 5312 |
Bound-variable hypothesis builder for a function with domain.
(Contributed by NM, 30-Jan-2004.)
|
β’ β²π₯πΉ
& β’ β²π₯π΄ β β’ β²π₯ πΉ Fn π΄ |
|
Theorem | fnfun 5313 |
A function with domain is a function. (Contributed by NM, 1-Aug-1994.)
|
β’ (πΉ Fn π΄ β Fun πΉ) |
|
Theorem | fnrel 5314 |
A function with domain is a relation. (Contributed by NM, 1-Aug-1994.)
|
β’ (πΉ Fn π΄ β Rel πΉ) |
|
Theorem | fndm 5315 |
The domain of a function. (Contributed by NM, 2-Aug-1994.)
|
β’ (πΉ Fn π΄ β dom πΉ = π΄) |
|
Theorem | funfni 5316 |
Inference to convert a function and domain antecedent. (Contributed by
NM, 22-Apr-2004.)
|
β’ ((Fun πΉ β§ π΅ β dom πΉ) β π) β β’ ((πΉ Fn π΄ β§ π΅ β π΄) β π) |
|
Theorem | fndmu 5317 |
A function has a unique domain. (Contributed by NM, 11-Aug-1994.)
|
β’ ((πΉ Fn π΄ β§ πΉ Fn π΅) β π΄ = π΅) |
|
Theorem | fnbr 5318 |
The first argument of binary relation on a function belongs to the
function's domain. (Contributed by NM, 7-May-2004.)
|
β’ ((πΉ Fn π΄ β§ π΅πΉπΆ) β π΅ β π΄) |
|
Theorem | fnop 5319 |
The first argument of an ordered pair in a function belongs to the
function's domain. (Contributed by NM, 8-Aug-1994.)
|
β’ ((πΉ Fn π΄ β§ β¨π΅, πΆβ© β πΉ) β π΅ β π΄) |
|
Theorem | fneu 5320* |
There is exactly one value of a function. (Contributed by NM,
22-Apr-2004.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ ((πΉ Fn π΄ β§ π΅ β π΄) β β!π¦ π΅πΉπ¦) |
|
Theorem | fneu2 5321* |
There is exactly one value of a function. (Contributed by NM,
7-Nov-1995.)
|
β’ ((πΉ Fn π΄ β§ π΅ β π΄) β β!π¦β¨π΅, π¦β© β πΉ) |
|
Theorem | fnun 5322 |
The union of two functions with disjoint domains. (Contributed by NM,
22-Sep-2004.)
|
β’ (((πΉ Fn π΄ β§ πΊ Fn π΅) β§ (π΄ β© π΅) = β
) β (πΉ βͺ πΊ) Fn (π΄ βͺ π΅)) |
|
Theorem | fnunsn 5323 |
Extension of a function with a new ordered pair. (Contributed by NM,
28-Sep-2013.) (Revised by Mario Carneiro, 30-Apr-2015.)
|
β’ (π β π β V) & β’ (π β π β V) & β’ (π β πΉ Fn π·)
& β’ πΊ = (πΉ βͺ {β¨π, πβ©}) & β’ πΈ = (π· βͺ {π}) & β’ (π β Β¬ π β π·) β β’ (π β πΊ Fn πΈ) |
|
Theorem | fnco 5324 |
Composition of two functions. (Contributed by NM, 22-May-2006.)
|
β’ ((πΉ Fn π΄ β§ πΊ Fn π΅ β§ ran πΊ β π΄) β (πΉ β πΊ) Fn π΅) |
|
Theorem | fnresdm 5325 |
A function does not change when restricted to its domain. (Contributed by
NM, 5-Sep-2004.)
|
β’ (πΉ Fn π΄ β (πΉ βΎ π΄) = πΉ) |
|
Theorem | fnresdisj 5326 |
A function restricted to a class disjoint with its domain is empty.
(Contributed by NM, 23-Sep-2004.)
|
β’ (πΉ Fn π΄ β ((π΄ β© π΅) = β
β (πΉ βΎ π΅) = β
)) |
|
Theorem | 2elresin 5327 |
Membership in two functions restricted by each other's domain.
(Contributed by NM, 8-Aug-1994.)
|
β’ ((πΉ Fn π΄ β§ πΊ Fn π΅) β ((β¨π₯, π¦β© β πΉ β§ β¨π₯, π§β© β πΊ) β (β¨π₯, π¦β© β (πΉ βΎ (π΄ β© π΅)) β§ β¨π₯, π§β© β (πΊ βΎ (π΄ β© π΅))))) |
|
Theorem | fnssresb 5328 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 10-Oct-2007.)
|
β’ (πΉ Fn π΄ β ((πΉ βΎ π΅) Fn π΅ β π΅ β π΄)) |
|
Theorem | fnssres 5329 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 2-Aug-1994.)
|
β’ ((πΉ Fn π΄ β§ π΅ β π΄) β (πΉ βΎ π΅) Fn π΅) |
|
Theorem | fnresin1 5330 |
Restriction of a function's domain with an intersection. (Contributed by
NM, 9-Aug-1994.)
|
β’ (πΉ Fn π΄ β (πΉ βΎ (π΄ β© π΅)) Fn (π΄ β© π΅)) |
|
Theorem | fnresin2 5331 |
Restriction of a function's domain with an intersection. (Contributed by
NM, 9-Aug-1994.)
|
β’ (πΉ Fn π΄ β (πΉ βΎ (π΅ β© π΄)) Fn (π΅ β© π΄)) |
|
Theorem | fnres 5332* |
An equivalence for functionality of a restriction. Compare dffun8 5244.
(Contributed by Mario Carneiro, 20-May-2015.)
|
β’ ((πΉ βΎ π΄) Fn π΄ β βπ₯ β π΄ β!π¦ π₯πΉπ¦) |
|
Theorem | fnresi 5333 |
Functionality and domain of restricted identity. (Contributed by NM,
27-Aug-2004.)
|
β’ ( I βΎ π΄) Fn π΄ |
|
Theorem | fnima 5334 |
The image of a function's domain is its range. (Contributed by NM,
4-Nov-2004.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ Fn π΄ β (πΉ β π΄) = ran πΉ) |
|
Theorem | fn0 5335 |
A function with empty domain is empty. (Contributed by NM, 15-Apr-1998.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ Fn β
β πΉ = β
) |
|
Theorem | fnimadisj 5336 |
A class that is disjoint with the domain of a function has an empty image
under the function. (Contributed by FL, 24-Jan-2007.)
|
β’ ((πΉ Fn π΄ β§ (π΄ β© πΆ) = β
) β (πΉ β πΆ) = β
) |
|
Theorem | fnimaeq0 5337 |
Images under a function never map nonempty sets to empty sets.
(Contributed by Stefan O'Rear, 21-Jan-2015.)
|
β’ ((πΉ Fn π΄ β§ π΅ β π΄) β ((πΉ β π΅) = β
β π΅ = β
)) |
|
Theorem | dfmpt3 5338 |
Alternate definition for the maps-to notation df-mpt 4066. (Contributed
by Mario Carneiro, 30-Dec-2016.)
|
β’ (π₯ β π΄ β¦ π΅) = βͺ
π₯ β π΄ ({π₯} Γ {π΅}) |
|
Theorem | fnopabg 5339* |
Functionality and domain of an ordered-pair class abstraction.
(Contributed by NM, 30-Jan-2004.) (Proof shortened by Mario Carneiro,
4-Dec-2016.)
|
β’ πΉ = {β¨π₯, π¦β© β£ (π₯ β π΄ β§ π)} β β’ (βπ₯ β π΄ β!π¦π β πΉ Fn π΄) |
|
Theorem | fnopab 5340* |
Functionality and domain of an ordered-pair class abstraction.
(Contributed by NM, 5-Mar-1996.)
|
β’ (π₯ β π΄ β β!π¦π)
& β’ πΉ = {β¨π₯, π¦β© β£ (π₯ β π΄ β§ π)} β β’ πΉ Fn π΄ |
|
Theorem | mptfng 5341* |
The maps-to notation defines a function with domain. (Contributed by
Scott Fenton, 21-Mar-2011.)
|
β’ πΉ = (π₯ β π΄ β¦ π΅) β β’ (βπ₯ β π΄ π΅ β V β πΉ Fn π΄) |
|
Theorem | fnmpt 5342* |
The maps-to notation defines a function with domain. (Contributed by
NM, 9-Apr-2013.)
|
β’ πΉ = (π₯ β π΄ β¦ π΅) β β’ (βπ₯ β π΄ π΅ β π β πΉ Fn π΄) |
|
Theorem | mpt0 5343 |
A mapping operation with empty domain. (Contributed by Mario Carneiro,
28-Dec-2014.)
|
β’ (π₯ β β
β¦ π΄) = β
|
|
Theorem | fnmpti 5344* |
Functionality and domain of an ordered-pair class abstraction.
(Contributed by NM, 29-Jan-2004.) (Revised by Mario Carneiro,
31-Aug-2015.)
|
β’ π΅ β V & β’ πΉ = (π₯ β π΄ β¦ π΅) β β’ πΉ Fn π΄ |
|
Theorem | dmmpti 5345* |
Domain of an ordered-pair class abstraction that specifies a function.
(Contributed by NM, 6-Sep-2005.) (Revised by Mario Carneiro,
31-Aug-2015.)
|
β’ π΅ β V & β’ πΉ = (π₯ β π΄ β¦ π΅) β β’ dom πΉ = π΄ |
|
Theorem | dmmptd 5346* |
The domain of the mapping operation, deduction form. (Contributed by
Glauco Siliprandi, 11-Dec-2019.)
|
β’ π΄ = (π₯ β π΅ β¦ πΆ)
& β’ ((π β§ π₯ β π΅) β πΆ β π) β β’ (π β dom π΄ = π΅) |
|
Theorem | mptun 5347 |
Union of mappings which are mutually compatible. (Contributed by Mario
Carneiro, 31-Aug-2015.)
|
β’ (π₯ β (π΄ βͺ π΅) β¦ πΆ) = ((π₯ β π΄ β¦ πΆ) βͺ (π₯ β π΅ β¦ πΆ)) |
|
Theorem | feq1 5348 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
β’ (πΉ = πΊ β (πΉ:π΄βΆπ΅ β πΊ:π΄βΆπ΅)) |
|
Theorem | feq2 5349 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
β’ (π΄ = π΅ β (πΉ:π΄βΆπΆ β πΉ:π΅βΆπΆ)) |
|
Theorem | feq3 5350 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
β’ (π΄ = π΅ β (πΉ:πΆβΆπ΄ β πΉ:πΆβΆπ΅)) |
|
Theorem | feq23 5351 |
Equality theorem for functions. (Contributed by FL, 14-Jul-2007.) (Proof
shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ ((π΄ = πΆ β§ π΅ = π·) β (πΉ:π΄βΆπ΅ β πΉ:πΆβΆπ·)) |
|
Theorem | feq1d 5352 |
Equality deduction for functions. (Contributed by NM, 19-Feb-2008.)
|
β’ (π β πΉ = πΊ) β β’ (π β (πΉ:π΄βΆπ΅ β πΊ:π΄βΆπ΅)) |
|
Theorem | feq2d 5353 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
β’ (π β π΄ = π΅) β β’ (π β (πΉ:π΄βΆπΆ β πΉ:π΅βΆπΆ)) |
|
Theorem | feq3d 5354 |
Equality deduction for functions. (Contributed by AV, 1-Jan-2020.)
|
β’ (π β π΄ = π΅) β β’ (π β (πΉ:πβΆπ΄ β πΉ:πβΆπ΅)) |
|
Theorem | feq12d 5355 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
β’ (π β πΉ = πΊ)
& β’ (π β π΄ = π΅) β β’ (π β (πΉ:π΄βΆπΆ β πΊ:π΅βΆπΆ)) |
|
Theorem | feq123d 5356 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
β’ (π β πΉ = πΊ)
& β’ (π β π΄ = π΅)
& β’ (π β πΆ = π·) β β’ (π β (πΉ:π΄βΆπΆ β πΊ:π΅βΆπ·)) |
|
Theorem | feq123 5357 |
Equality theorem for functions. (Contributed by FL, 16-Nov-2008.)
|
β’ ((πΉ = πΊ β§ π΄ = πΆ β§ π΅ = π·) β (πΉ:π΄βΆπ΅ β πΊ:πΆβΆπ·)) |
|
Theorem | feq1i 5358 |
Equality inference for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
β’ πΉ = πΊ β β’ (πΉ:π΄βΆπ΅ β πΊ:π΄βΆπ΅) |
|
Theorem | feq2i 5359 |
Equality inference for functions. (Contributed by NM, 5-Sep-2011.)
|
β’ π΄ = π΅ β β’ (πΉ:π΄βΆπΆ β πΉ:π΅βΆπΆ) |
|
Theorem | feq23i 5360 |
Equality inference for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
β’ π΄ = πΆ
& β’ π΅ = π· β β’ (πΉ:π΄βΆπ΅ β πΉ:πΆβΆπ·) |
|
Theorem | feq23d 5361 |
Equality deduction for functions. (Contributed by NM, 8-Jun-2013.)
|
β’ (π β π΄ = πΆ)
& β’ (π β π΅ = π·) β β’ (π β (πΉ:π΄βΆπ΅ β πΉ:πΆβΆπ·)) |
|
Theorem | nff 5362 |
Bound-variable hypothesis builder for a mapping. (Contributed by NM,
29-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.)
|
β’ β²π₯πΉ
& β’ β²π₯π΄
& β’ β²π₯π΅ β β’ β²π₯ πΉ:π΄βΆπ΅ |
|
Theorem | sbcfng 5363* |
Distribute proper substitution through the function predicate with a
domain. (Contributed by Alexander van der Vekens, 15-Jul-2018.)
|
β’ (π β π β ([π / π₯]πΉ Fn π΄ β β¦π / π₯β¦πΉ Fn β¦π / π₯β¦π΄)) |
|
Theorem | sbcfg 5364* |
Distribute proper substitution through the function predicate with
domain and codomain. (Contributed by Alexander van der Vekens,
15-Jul-2018.)
|
β’ (π β π β ([π / π₯]πΉ:π΄βΆπ΅ β β¦π / π₯β¦πΉ:β¦π / π₯β¦π΄βΆβ¦π / π₯β¦π΅)) |
|
Theorem | ffn 5365 |
A mapping is a function. (Contributed by NM, 2-Aug-1994.)
|
β’ (πΉ:π΄βΆπ΅ β πΉ Fn π΄) |
|
Theorem | ffnd 5366 |
A mapping is a function with domain, deduction form. (Contributed by
Glauco Siliprandi, 17-Aug-2020.)
|
β’ (π β πΉ:π΄βΆπ΅) β β’ (π β πΉ Fn π΄) |
|
Theorem | dffn2 5367 |
Any function is a mapping into V. (Contributed by NM,
31-Oct-1995.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ Fn π΄ β πΉ:π΄βΆV) |
|
Theorem | ffun 5368 |
A mapping is a function. (Contributed by NM, 3-Aug-1994.)
|
β’ (πΉ:π΄βΆπ΅ β Fun πΉ) |
|
Theorem | ffund 5369 |
A mapping is a function, deduction version. (Contributed by Glauco
Siliprandi, 3-Mar-2021.)
|
β’ (π β πΉ:π΄βΆπ΅) β β’ (π β Fun πΉ) |
|
Theorem | frel 5370 |
A mapping is a relation. (Contributed by NM, 3-Aug-1994.)
|
β’ (πΉ:π΄βΆπ΅ β Rel πΉ) |
|
Theorem | fdm 5371 |
The domain of a mapping. (Contributed by NM, 2-Aug-1994.)
|
β’ (πΉ:π΄βΆπ΅ β dom πΉ = π΄) |
|
Theorem | fdmd 5372 |
Deduction form of fdm 5371. The domain of a mapping. (Contributed by
Glauco Siliprandi, 26-Jun-2021.)
|
β’ (π β πΉ:π΄βΆπ΅) β β’ (π β dom πΉ = π΄) |
|
Theorem | fdmi 5373 |
The domain of a mapping. (Contributed by NM, 28-Jul-2008.)
|
β’ πΉ:π΄βΆπ΅ β β’ dom πΉ = π΄ |
|
Theorem | frn 5374 |
The range of a mapping. (Contributed by NM, 3-Aug-1994.)
|
β’ (πΉ:π΄βΆπ΅ β ran πΉ β π΅) |
|
Theorem | frnd 5375 |
Deduction form of frn 5374. The range of a mapping. (Contributed by
Glauco Siliprandi, 26-Jun-2021.)
|
β’ (π β πΉ:π΄βΆπ΅) β β’ (π β ran πΉ β π΅) |
|
Theorem | dffn3 5376 |
A function maps to its range. (Contributed by NM, 1-Sep-1999.)
|
β’ (πΉ Fn π΄ β πΉ:π΄βΆran πΉ) |
|
Theorem | fss 5377 |
Expanding the codomain of a mapping. (Contributed by NM, 10-May-1998.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ ((πΉ:π΄βΆπ΅ β§ π΅ β πΆ) β πΉ:π΄βΆπΆ) |
|
Theorem | fssd 5378 |
Expanding the codomain of a mapping, deduction form. (Contributed by
Glauco Siliprandi, 11-Dec-2019.)
|
β’ (π β πΉ:π΄βΆπ΅)
& β’ (π β π΅ β πΆ) β β’ (π β πΉ:π΄βΆπΆ) |
|
Theorem | fssdmd 5379 |
Expressing that a class is a subclass of the domain of a function
expressed in maps-to notation, deduction form. (Contributed by AV,
21-Aug-2022.)
|
β’ (π β πΉ:π΄βΆπ΅)
& β’ (π β π· β dom πΉ) β β’ (π β π· β π΄) |
|
Theorem | fssdm 5380 |
Expressing that a class is a subclass of the domain of a function
expressed in maps-to notation, semi-deduction form. (Contributed by AV,
21-Aug-2022.)
|
β’ π· β dom πΉ
& β’ (π β πΉ:π΄βΆπ΅) β β’ (π β π· β π΄) |
|
Theorem | fco 5381 |
Composition of two mappings. (Contributed by NM, 29-Aug-1999.) (Proof
shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ ((πΉ:π΅βΆπΆ β§ πΊ:π΄βΆπ΅) β (πΉ β πΊ):π΄βΆπΆ) |
|
Theorem | fco2 5382 |
Functionality of a composition with weakened out of domain condition on
the first argument. (Contributed by Stefan O'Rear, 11-Mar-2015.)
|
β’ (((πΉ βΎ π΅):π΅βΆπΆ β§ πΊ:π΄βΆπ΅) β (πΉ β πΊ):π΄βΆπΆ) |
|
Theorem | fssxp 5383 |
A mapping is a class of ordered pairs. (Contributed by NM, 3-Aug-1994.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ:π΄βΆπ΅ β πΉ β (π΄ Γ π΅)) |
|
Theorem | fex2 5384 |
A function with bounded domain and codomain is a set. This version is
proven without the Axiom of Replacement. (Contributed by Mario Carneiro,
24-Jun-2015.)
|
β’ ((πΉ:π΄βΆπ΅ β§ π΄ β π β§ π΅ β π) β πΉ β V) |
|
Theorem | funssxp 5385 |
Two ways of specifying a partial function from π΄ to π΅.
(Contributed by NM, 13-Nov-2007.)
|
β’ ((Fun πΉ β§ πΉ β (π΄ Γ π΅)) β (πΉ:dom πΉβΆπ΅ β§ dom πΉ β π΄)) |
|
Theorem | ffdm 5386 |
A mapping is a partial function. (Contributed by NM, 25-Nov-2007.)
|
β’ (πΉ:π΄βΆπ΅ β (πΉ:dom πΉβΆπ΅ β§ dom πΉ β π΄)) |
|
Theorem | opelf 5387 |
The members of an ordered pair element of a mapping belong to the
mapping's domain and codomain. (Contributed by NM, 10-Dec-2003.)
(Revised by Mario Carneiro, 26-Apr-2015.)
|
β’ ((πΉ:π΄βΆπ΅ β§ β¨πΆ, π·β© β πΉ) β (πΆ β π΄ β§ π· β π΅)) |
|
Theorem | fun 5388 |
The union of two functions with disjoint domains. (Contributed by NM,
22-Sep-2004.)
|
β’ (((πΉ:π΄βΆπΆ β§ πΊ:π΅βΆπ·) β§ (π΄ β© π΅) = β
) β (πΉ βͺ πΊ):(π΄ βͺ π΅)βΆ(πΆ βͺ π·)) |
|
Theorem | fun2 5389 |
The union of two functions with disjoint domains. (Contributed by Mario
Carneiro, 12-Mar-2015.)
|
β’ (((πΉ:π΄βΆπΆ β§ πΊ:π΅βΆπΆ) β§ (π΄ β© π΅) = β
) β (πΉ βͺ πΊ):(π΄ βͺ π΅)βΆπΆ) |
|
Theorem | fnfco 5390 |
Composition of two functions. (Contributed by NM, 22-May-2006.)
|
β’ ((πΉ Fn π΄ β§ πΊ:π΅βΆπ΄) β (πΉ β πΊ) Fn π΅) |
|
Theorem | fssres 5391 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 23-Sep-2004.)
|
β’ ((πΉ:π΄βΆπ΅ β§ πΆ β π΄) β (πΉ βΎ πΆ):πΆβΆπ΅) |
|
Theorem | fssresd 5392 |
Restriction of a function with a subclass of its domain, deduction form.
(Contributed by Glauco Siliprandi, 11-Dec-2019.)
|
β’ (π β πΉ:π΄βΆπ΅)
& β’ (π β πΆ β π΄) β β’ (π β (πΉ βΎ πΆ):πΆβΆπ΅) |
|
Theorem | fssres2 5393 |
Restriction of a restricted function with a subclass of its domain.
(Contributed by NM, 21-Jul-2005.)
|
β’ (((πΉ βΎ π΄):π΄βΆπ΅ β§ πΆ β π΄) β (πΉ βΎ πΆ):πΆβΆπ΅) |
|
Theorem | fresin 5394 |
An identity for the mapping relationship under restriction. (Contributed
by Scott Fenton, 4-Sep-2011.) (Proof shortened by Mario Carneiro,
26-May-2016.)
|
β’ (πΉ:π΄βΆπ΅ β (πΉ βΎ π):(π΄ β© π)βΆπ΅) |
|
Theorem | resasplitss 5395 |
If two functions agree on their common domain, their union contains a
union of three functions with pairwise disjoint domains. If we assumed
the law of the excluded middle, this would be equality rather than subset.
(Contributed by Jim Kingdon, 28-Dec-2018.)
|
β’ ((πΉ Fn π΄ β§ πΊ Fn π΅ β§ (πΉ βΎ (π΄ β© π΅)) = (πΊ βΎ (π΄ β© π΅))) β ((πΉ βΎ (π΄ β© π΅)) βͺ ((πΉ βΎ (π΄ β π΅)) βͺ (πΊ βΎ (π΅ β π΄)))) β (πΉ βͺ πΊ)) |
|
Theorem | fcoi1 5396 |
Composition of a mapping and restricted identity. (Contributed by NM,
13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ:π΄βΆπ΅ β (πΉ β ( I βΎ π΄)) = πΉ) |
|
Theorem | fcoi2 5397 |
Composition of restricted identity and a mapping. (Contributed by NM,
13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
β’ (πΉ:π΄βΆπ΅ β (( I βΎ π΅) β πΉ) = πΉ) |
|
Theorem | feu 5398* |
There is exactly one value of a function in its codomain. (Contributed
by NM, 10-Dec-2003.)
|
β’ ((πΉ:π΄βΆπ΅ β§ πΆ β π΄) β β!π¦ β π΅ β¨πΆ, π¦β© β πΉ) |
|
Theorem | fcnvres 5399 |
The converse of a restriction of a function. (Contributed by NM,
26-Mar-1998.)
|
β’ (πΉ:π΄βΆπ΅ β β‘(πΉ βΎ π΄) = (β‘πΉ βΎ π΅)) |
|
Theorem | fimacnvdisj 5400 |
The preimage of a class disjoint with a mapping's codomain is empty.
(Contributed by FL, 24-Jan-2007.)
|
β’ ((πΉ:π΄βΆπ΅ β§ (π΅ β© πΆ) = β
) β (β‘πΉ β πΆ) = β
) |