Theorem List for Intuitionistic Logic Explorer - 5301-5400 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Theorem | fnop 5301 |
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 5302* |
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 5303* |
There is exactly one value of a function. (Contributed by NM,
7-Nov-1995.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → ∃!𝑦〈𝐵, 𝑦〉 ∈ 𝐹) |
|
Theorem | fnun 5304 |
The union of two functions with disjoint domains. (Contributed by NM,
22-Sep-2004.)
|
⊢ (((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐹 ∪ 𝐺) Fn (𝐴 ∪ 𝐵)) |
|
Theorem | fnunsn 5305 |
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 5306 |
Composition of two functions. (Contributed by NM, 22-May-2006.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵 ∧ ran 𝐺 ⊆ 𝐴) → (𝐹 ∘ 𝐺) Fn 𝐵) |
|
Theorem | fnresdm 5307 |
A function does not change when restricted to its domain. (Contributed by
NM, 5-Sep-2004.)
|
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ 𝐴) = 𝐹) |
|
Theorem | fnresdisj 5308 |
A function restricted to a class disjoint with its domain is empty.
(Contributed by NM, 23-Sep-2004.)
|
⊢ (𝐹 Fn 𝐴 → ((𝐴 ∩ 𝐵) = ∅ ↔ (𝐹 ↾ 𝐵) = ∅)) |
|
Theorem | 2elresin 5309 |
Membership in two functions restricted by each other's domain.
(Contributed by NM, 8-Aug-1994.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) → ((〈𝑥, 𝑦〉 ∈ 𝐹 ∧ 〈𝑥, 𝑧〉 ∈ 𝐺) ↔ (〈𝑥, 𝑦〉 ∈ (𝐹 ↾ (𝐴 ∩ 𝐵)) ∧ 〈𝑥, 𝑧〉 ∈ (𝐺 ↾ (𝐴 ∩ 𝐵))))) |
|
Theorem | fnssresb 5310 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 10-Oct-2007.)
|
⊢ (𝐹 Fn 𝐴 → ((𝐹 ↾ 𝐵) Fn 𝐵 ↔ 𝐵 ⊆ 𝐴)) |
|
Theorem | fnssres 5311 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 2-Aug-1994.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ⊆ 𝐴) → (𝐹 ↾ 𝐵) Fn 𝐵) |
|
Theorem | fnresin1 5312 |
Restriction of a function's domain with an intersection. (Contributed by
NM, 9-Aug-1994.)
|
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐴 ∩ 𝐵)) Fn (𝐴 ∩ 𝐵)) |
|
Theorem | fnresin2 5313 |
Restriction of a function's domain with an intersection. (Contributed by
NM, 9-Aug-1994.)
|
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐵 ∩ 𝐴)) Fn (𝐵 ∩ 𝐴)) |
|
Theorem | fnres 5314* |
An equivalence for functionality of a restriction. Compare dffun8 5226.
(Contributed by Mario Carneiro, 20-May-2015.)
|
⊢ ((𝐹 ↾ 𝐴) Fn 𝐴 ↔ ∀𝑥 ∈ 𝐴 ∃!𝑦 𝑥𝐹𝑦) |
|
Theorem | fnresi 5315 |
Functionality and domain of restricted identity. (Contributed by NM,
27-Aug-2004.)
|
⊢ ( I ↾ 𝐴) Fn 𝐴 |
|
Theorem | fnima 5316 |
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 5317 |
A function with empty domain is empty. (Contributed by NM, 15-Apr-1998.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (𝐹 Fn ∅ ↔ 𝐹 = ∅) |
|
Theorem | fnimadisj 5318 |
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 5319 |
Images under a function never map nonempty sets to empty sets.
(Contributed by Stefan O'Rear, 21-Jan-2015.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ⊆ 𝐴) → ((𝐹 “ 𝐵) = ∅ ↔ 𝐵 = ∅)) |
|
Theorem | dfmpt3 5320 |
Alternate definition for the maps-to notation df-mpt 4052. (Contributed
by Mario Carneiro, 30-Dec-2016.)
|
⊢ (𝑥 ∈ 𝐴 ↦ 𝐵) = ∪
𝑥 ∈ 𝐴 ({𝑥} × {𝐵}) |
|
Theorem | fnopabg 5321* |
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 5322* |
Functionality and domain of an ordered-pair class abstraction.
(Contributed by NM, 5-Mar-1996.)
|
⊢ (𝑥 ∈ 𝐴 → ∃!𝑦𝜑)
& ⊢ 𝐹 = {〈𝑥, 𝑦〉 ∣ (𝑥 ∈ 𝐴 ∧ 𝜑)} ⇒ ⊢ 𝐹 Fn 𝐴 |
|
Theorem | mptfng 5323* |
The maps-to notation defines a function with domain. (Contributed by
Scott Fenton, 21-Mar-2011.)
|
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ V ↔ 𝐹 Fn 𝐴) |
|
Theorem | fnmpt 5324* |
The maps-to notation defines a function with domain. (Contributed by
NM, 9-Apr-2013.)
|
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ 𝑉 → 𝐹 Fn 𝐴) |
|
Theorem | mpt0 5325 |
A mapping operation with empty domain. (Contributed by Mario Carneiro,
28-Dec-2014.)
|
⊢ (𝑥 ∈ ∅ ↦ 𝐴) = ∅ |
|
Theorem | fnmpti 5326* |
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 5327* |
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 5328* |
The domain of the mapping operation, deduction form. (Contributed by
Glauco Siliprandi, 11-Dec-2019.)
|
⊢ 𝐴 = (𝑥 ∈ 𝐵 ↦ 𝐶)
& ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐵) → 𝐶 ∈ 𝑉) ⇒ ⊢ (𝜑 → dom 𝐴 = 𝐵) |
|
Theorem | mptun 5329 |
Union of mappings which are mutually compatible. (Contributed by Mario
Carneiro, 31-Aug-2015.)
|
⊢ (𝑥 ∈ (𝐴 ∪ 𝐵) ↦ 𝐶) = ((𝑥 ∈ 𝐴 ↦ 𝐶) ∪ (𝑥 ∈ 𝐵 ↦ 𝐶)) |
|
Theorem | feq1 5330 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
⊢ (𝐹 = 𝐺 → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵)) |
|
Theorem | feq2 5331 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
⊢ (𝐴 = 𝐵 → (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶)) |
|
Theorem | feq3 5332 |
Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
|
⊢ (𝐴 = 𝐵 → (𝐹:𝐶⟶𝐴 ↔ 𝐹:𝐶⟶𝐵)) |
|
Theorem | feq23 5333 |
Equality theorem for functions. (Contributed by FL, 14-Jul-2007.) (Proof
shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ ((𝐴 = 𝐶 ∧ 𝐵 = 𝐷) → (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷)) |
|
Theorem | feq1d 5334 |
Equality deduction for functions. (Contributed by NM, 19-Feb-2008.)
|
⊢ (𝜑 → 𝐹 = 𝐺) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵)) |
|
Theorem | feq2d 5335 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶)) |
|
Theorem | feq3d 5336 |
Equality deduction for functions. (Contributed by AV, 1-Jan-2020.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝑋⟶𝐴 ↔ 𝐹:𝑋⟶𝐵)) |
|
Theorem | feq12d 5337 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
⊢ (𝜑 → 𝐹 = 𝐺)
& ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐺:𝐵⟶𝐶)) |
|
Theorem | feq123d 5338 |
Equality deduction for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
⊢ (𝜑 → 𝐹 = 𝐺)
& ⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐺:𝐵⟶𝐷)) |
|
Theorem | feq123 5339 |
Equality theorem for functions. (Contributed by FL, 16-Nov-2008.)
|
⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐶 ∧ 𝐵 = 𝐷) → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐶⟶𝐷)) |
|
Theorem | feq1i 5340 |
Equality inference for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
⊢ 𝐹 = 𝐺 ⇒ ⊢ (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵) |
|
Theorem | feq2i 5341 |
Equality inference for functions. (Contributed by NM, 5-Sep-2011.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶) |
|
Theorem | feq23i 5342 |
Equality inference for functions. (Contributed by Paul Chapman,
22-Jun-2011.)
|
⊢ 𝐴 = 𝐶
& ⊢ 𝐵 = 𝐷 ⇒ ⊢ (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷) |
|
Theorem | feq23d 5343 |
Equality deduction for functions. (Contributed by NM, 8-Jun-2013.)
|
⊢ (𝜑 → 𝐴 = 𝐶)
& ⊢ (𝜑 → 𝐵 = 𝐷) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷)) |
|
Theorem | nff 5344 |
Bound-variable hypothesis builder for a mapping. (Contributed by NM,
29-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.)
|
⊢ Ⅎ𝑥𝐹
& ⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥 𝐹:𝐴⟶𝐵 |
|
Theorem | sbcfng 5345* |
Distribute proper substitution through the function predicate with a
domain. (Contributed by Alexander van der Vekens, 15-Jul-2018.)
|
⊢ (𝑋 ∈ 𝑉 → ([𝑋 / 𝑥]𝐹 Fn 𝐴 ↔ ⦋𝑋 / 𝑥⦌𝐹 Fn ⦋𝑋 / 𝑥⦌𝐴)) |
|
Theorem | sbcfg 5346* |
Distribute proper substitution through the function predicate with
domain and codomain. (Contributed by Alexander van der Vekens,
15-Jul-2018.)
|
⊢ (𝑋 ∈ 𝑉 → ([𝑋 / 𝑥]𝐹:𝐴⟶𝐵 ↔ ⦋𝑋 / 𝑥⦌𝐹:⦋𝑋 / 𝑥⦌𝐴⟶⦋𝑋 / 𝑥⦌𝐵)) |
|
Theorem | ffn 5347 |
A mapping is a function. (Contributed by NM, 2-Aug-1994.)
|
⊢ (𝐹:𝐴⟶𝐵 → 𝐹 Fn 𝐴) |
|
Theorem | ffnd 5348 |
A mapping is a function with domain, deduction form. (Contributed by
Glauco Siliprandi, 17-Aug-2020.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → 𝐹 Fn 𝐴) |
|
Theorem | dffn2 5349 |
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 5350 |
A mapping is a function. (Contributed by NM, 3-Aug-1994.)
|
⊢ (𝐹:𝐴⟶𝐵 → Fun 𝐹) |
|
Theorem | ffund 5351 |
A mapping is a function, deduction version. (Contributed by Glauco
Siliprandi, 3-Mar-2021.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → Fun 𝐹) |
|
Theorem | frel 5352 |
A mapping is a relation. (Contributed by NM, 3-Aug-1994.)
|
⊢ (𝐹:𝐴⟶𝐵 → Rel 𝐹) |
|
Theorem | fdm 5353 |
The domain of a mapping. (Contributed by NM, 2-Aug-1994.)
|
⊢ (𝐹:𝐴⟶𝐵 → dom 𝐹 = 𝐴) |
|
Theorem | fdmd 5354 |
Deduction form of fdm 5353. The domain of a mapping. (Contributed by
Glauco Siliprandi, 26-Jun-2021.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → dom 𝐹 = 𝐴) |
|
Theorem | fdmi 5355 |
The domain of a mapping. (Contributed by NM, 28-Jul-2008.)
|
⊢ 𝐹:𝐴⟶𝐵 ⇒ ⊢ dom 𝐹 = 𝐴 |
|
Theorem | frn 5356 |
The range of a mapping. (Contributed by NM, 3-Aug-1994.)
|
⊢ (𝐹:𝐴⟶𝐵 → ran 𝐹 ⊆ 𝐵) |
|
Theorem | frnd 5357 |
Deduction form of frn 5356. The range of a mapping. (Contributed by
Glauco Siliprandi, 26-Jun-2021.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → ran 𝐹 ⊆ 𝐵) |
|
Theorem | dffn3 5358 |
A function maps to its range. (Contributed by NM, 1-Sep-1999.)
|
⊢ (𝐹 Fn 𝐴 ↔ 𝐹:𝐴⟶ran 𝐹) |
|
Theorem | fss 5359 |
Expanding the codomain of a mapping. (Contributed by NM, 10-May-1998.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ ((𝐹:𝐴⟶𝐵 ∧ 𝐵 ⊆ 𝐶) → 𝐹:𝐴⟶𝐶) |
|
Theorem | fssd 5360 |
Expanding the codomain of a mapping, deduction form. (Contributed by
Glauco Siliprandi, 11-Dec-2019.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐹:𝐴⟶𝐶) |
|
Theorem | fssdmd 5361 |
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 5362 |
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 5363 |
Composition of two mappings. (Contributed by NM, 29-Aug-1999.) (Proof
shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ ((𝐹:𝐵⟶𝐶 ∧ 𝐺:𝐴⟶𝐵) → (𝐹 ∘ 𝐺):𝐴⟶𝐶) |
|
Theorem | fco2 5364 |
Functionality of a composition with weakened out of domain condition on
the first argument. (Contributed by Stefan O'Rear, 11-Mar-2015.)
|
⊢ (((𝐹 ↾ 𝐵):𝐵⟶𝐶 ∧ 𝐺:𝐴⟶𝐵) → (𝐹 ∘ 𝐺):𝐴⟶𝐶) |
|
Theorem | fssxp 5365 |
A mapping is a class of ordered pairs. (Contributed by NM, 3-Aug-1994.)
(Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (𝐹:𝐴⟶𝐵 → 𝐹 ⊆ (𝐴 × 𝐵)) |
|
Theorem | fex2 5366 |
A function with bounded domain and range is a set. This version is proven
without the Axiom of Replacement. (Contributed by Mario Carneiro,
24-Jun-2015.)
|
⊢ ((𝐹:𝐴⟶𝐵 ∧ 𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → 𝐹 ∈ V) |
|
Theorem | funssxp 5367 |
Two ways of specifying a partial function from 𝐴 to 𝐵.
(Contributed by NM, 13-Nov-2007.)
|
⊢ ((Fun 𝐹 ∧ 𝐹 ⊆ (𝐴 × 𝐵)) ↔ (𝐹:dom 𝐹⟶𝐵 ∧ dom 𝐹 ⊆ 𝐴)) |
|
Theorem | ffdm 5368 |
A mapping is a partial function. (Contributed by NM, 25-Nov-2007.)
|
⊢ (𝐹:𝐴⟶𝐵 → (𝐹:dom 𝐹⟶𝐵 ∧ dom 𝐹 ⊆ 𝐴)) |
|
Theorem | opelf 5369 |
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 5370 |
The union of two functions with disjoint domains. (Contributed by NM,
22-Sep-2004.)
|
⊢ (((𝐹:𝐴⟶𝐶 ∧ 𝐺:𝐵⟶𝐷) ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐹 ∪ 𝐺):(𝐴 ∪ 𝐵)⟶(𝐶 ∪ 𝐷)) |
|
Theorem | fun2 5371 |
The union of two functions with disjoint domains. (Contributed by Mario
Carneiro, 12-Mar-2015.)
|
⊢ (((𝐹:𝐴⟶𝐶 ∧ 𝐺:𝐵⟶𝐶) ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐹 ∪ 𝐺):(𝐴 ∪ 𝐵)⟶𝐶) |
|
Theorem | fnfco 5372 |
Composition of two functions. (Contributed by NM, 22-May-2006.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐺:𝐵⟶𝐴) → (𝐹 ∘ 𝐺) Fn 𝐵) |
|
Theorem | fssres 5373 |
Restriction of a function with a subclass of its domain. (Contributed by
NM, 23-Sep-2004.)
|
⊢ ((𝐹:𝐴⟶𝐵 ∧ 𝐶 ⊆ 𝐴) → (𝐹 ↾ 𝐶):𝐶⟶𝐵) |
|
Theorem | fssresd 5374 |
Restriction of a function with a subclass of its domain, deduction form.
(Contributed by Glauco Siliprandi, 11-Dec-2019.)
|
⊢ (𝜑 → 𝐹:𝐴⟶𝐵)
& ⊢ (𝜑 → 𝐶 ⊆ 𝐴) ⇒ ⊢ (𝜑 → (𝐹 ↾ 𝐶):𝐶⟶𝐵) |
|
Theorem | fssres2 5375 |
Restriction of a restricted function with a subclass of its domain.
(Contributed by NM, 21-Jul-2005.)
|
⊢ (((𝐹 ↾ 𝐴):𝐴⟶𝐵 ∧ 𝐶 ⊆ 𝐴) → (𝐹 ↾ 𝐶):𝐶⟶𝐵) |
|
Theorem | fresin 5376 |
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 5377 |
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 5378 |
Composition of a mapping and restricted identity. (Contributed by NM,
13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (𝐹:𝐴⟶𝐵 → (𝐹 ∘ ( I ↾ 𝐴)) = 𝐹) |
|
Theorem | fcoi2 5379 |
Composition of restricted identity and a mapping. (Contributed by NM,
13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (𝐹:𝐴⟶𝐵 → (( I ↾ 𝐵) ∘ 𝐹) = 𝐹) |
|
Theorem | feu 5380* |
There is exactly one value of a function in its codomain. (Contributed
by NM, 10-Dec-2003.)
|
⊢ ((𝐹:𝐴⟶𝐵 ∧ 𝐶 ∈ 𝐴) → ∃!𝑦 ∈ 𝐵 〈𝐶, 𝑦〉 ∈ 𝐹) |
|
Theorem | fcnvres 5381 |
The converse of a restriction of a function. (Contributed by NM,
26-Mar-1998.)
|
⊢ (𝐹:𝐴⟶𝐵 → ◡(𝐹 ↾ 𝐴) = (◡𝐹 ↾ 𝐵)) |
|
Theorem | fimacnvdisj 5382 |
The preimage of a class disjoint with a mapping's codomain is empty.
(Contributed by FL, 24-Jan-2007.)
|
⊢ ((𝐹:𝐴⟶𝐵 ∧ (𝐵 ∩ 𝐶) = ∅) → (◡𝐹 “ 𝐶) = ∅) |
|
Theorem | fintm 5383* |
Function into an intersection. (Contributed by Jim Kingdon,
28-Dec-2018.)
|
⊢ ∃𝑥 𝑥 ∈ 𝐵 ⇒ ⊢ (𝐹:𝐴⟶∩ 𝐵 ↔ ∀𝑥 ∈ 𝐵 𝐹:𝐴⟶𝑥) |
|
Theorem | fin 5384 |
Mapping into an intersection. (Contributed by NM, 14-Sep-1999.) (Proof
shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (𝐹:𝐴⟶(𝐵 ∩ 𝐶) ↔ (𝐹:𝐴⟶𝐵 ∧ 𝐹:𝐴⟶𝐶)) |
|
Theorem | fabexg 5385* |
Existence of a set of functions. (Contributed by Paul Chapman,
25-Feb-2008.)
|
⊢ 𝐹 = {𝑥 ∣ (𝑥:𝐴⟶𝐵 ∧ 𝜑)} ⇒ ⊢ ((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) → 𝐹 ∈ V) |
|
Theorem | fabex 5386* |
Existence of a set of functions. (Contributed by NM, 3-Dec-2007.)
|
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐹 = {𝑥 ∣ (𝑥:𝐴⟶𝐵 ∧ 𝜑)} ⇒ ⊢ 𝐹 ∈ V |
|
Theorem | dmfex 5387 |
If a mapping is a set, its domain is a set. (Contributed by NM,
27-Aug-2006.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ ((𝐹 ∈ 𝐶 ∧ 𝐹:𝐴⟶𝐵) → 𝐴 ∈ V) |
|
Theorem | f0 5388 |
The empty function. (Contributed by NM, 14-Aug-1999.)
|
⊢ ∅:∅⟶𝐴 |
|
Theorem | f00 5389 |
A class is a function with empty codomain iff it and its domain are empty.
(Contributed by NM, 10-Dec-2003.)
|
⊢ (𝐹:𝐴⟶∅ ↔ (𝐹 = ∅ ∧ 𝐴 = ∅)) |
|
Theorem | f0bi 5390 |
A function with empty domain is empty. (Contributed by Alexander van der
Vekens, 30-Jun-2018.)
|
⊢ (𝐹:∅⟶𝑋 ↔ 𝐹 = ∅) |
|
Theorem | f0dom0 5391 |
A function is empty iff it has an empty domain. (Contributed by AV,
10-Feb-2019.)
|
⊢ (𝐹:𝑋⟶𝑌 → (𝑋 = ∅ ↔ 𝐹 = ∅)) |
|
Theorem | f0rn0 5392* |
If there is no element in the range of a function, its domain must be
empty. (Contributed by Alexander van der Vekens, 12-Jul-2018.)
|
⊢ ((𝐸:𝑋⟶𝑌 ∧ ¬ ∃𝑦 ∈ 𝑌 𝑦 ∈ ran 𝐸) → 𝑋 = ∅) |
|
Theorem | fconst 5393 |
A cross product with a singleton is a constant function. (Contributed
by NM, 14-Aug-1999.) (Proof shortened by Andrew Salmon,
17-Sep-2011.)
|
⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 × {𝐵}):𝐴⟶{𝐵} |
|
Theorem | fconstg 5394 |
A cross product with a singleton is a constant function. (Contributed
by NM, 19-Oct-2004.)
|
⊢ (𝐵 ∈ 𝑉 → (𝐴 × {𝐵}):𝐴⟶{𝐵}) |
|
Theorem | fnconstg 5395 |
A cross product with a singleton is a constant function. (Contributed by
NM, 24-Jul-2014.)
|
⊢ (𝐵 ∈ 𝑉 → (𝐴 × {𝐵}) Fn 𝐴) |
|
Theorem | fconst6g 5396 |
Constant function with loose range. (Contributed by Stefan O'Rear,
1-Feb-2015.)
|
⊢ (𝐵 ∈ 𝐶 → (𝐴 × {𝐵}):𝐴⟶𝐶) |
|
Theorem | fconst6 5397 |
A constant function as a mapping. (Contributed by Jeff Madsen,
30-Nov-2009.) (Revised by Mario Carneiro, 22-Apr-2015.)
|
⊢ 𝐵 ∈ 𝐶 ⇒ ⊢ (𝐴 × {𝐵}):𝐴⟶𝐶 |
|
Theorem | f1eq1 5398 |
Equality theorem for one-to-one functions. (Contributed by NM,
10-Feb-1997.)
|
⊢ (𝐹 = 𝐺 → (𝐹:𝐴–1-1→𝐵 ↔ 𝐺:𝐴–1-1→𝐵)) |
|
Theorem | f1eq2 5399 |
Equality theorem for one-to-one functions. (Contributed by NM,
10-Feb-1997.)
|
⊢ (𝐴 = 𝐵 → (𝐹:𝐴–1-1→𝐶 ↔ 𝐹:𝐵–1-1→𝐶)) |
|
Theorem | f1eq3 5400 |
Equality theorem for one-to-one functions. (Contributed by NM,
10-Feb-1997.)
|
⊢ (𝐴 = 𝐵 → (𝐹:𝐶–1-1→𝐴 ↔ 𝐹:𝐶–1-1→𝐵)) |