| Metamath
Proof Explorer Theorem List (p. 67 of 494) | < Previous Next > | |
| Bad symbols? Try the
GIF version. |
||
|
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
||
| Color key: | (1-30937) |
(30938-32460) |
(32461-49324) |
| Type | Label | Description |
|---|---|---|
| Statement | ||
| Theorem | funopab 6601* | A class of ordered pairs is a function when there is at most one second member for each pair. (Contributed by NM, 16-May-1995.) |
| ⊢ (Fun {〈𝑥, 𝑦〉 ∣ 𝜑} ↔ ∀𝑥∃*𝑦𝜑) | ||
| Theorem | funopabeq 6602* | A class of ordered pairs of values is a function. (Contributed by NM, 14-Nov-1995.) |
| ⊢ Fun {〈𝑥, 𝑦〉 ∣ 𝑦 = 𝐴} | ||
| Theorem | funopab4 6603* | A class of ordered pairs of values in the form used by df-mpt 5226 is a function. (Contributed by NM, 17-Feb-2013.) |
| ⊢ Fun {〈𝑥, 𝑦〉 ∣ (𝜑 ∧ 𝑦 = 𝐴)} | ||
| Theorem | funmpt 6604 | A function in maps-to notation is a function. (Contributed by Mario Carneiro, 13-Jan-2013.) |
| ⊢ Fun (𝑥 ∈ 𝐴 ↦ 𝐵) | ||
| Theorem | funmpt2 6605 | Functionality of a class given by a maps-to notation. (Contributed by FL, 17-Feb-2008.) (Revised by Mario Carneiro, 31-May-2014.) |
| ⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ Fun 𝐹 | ||
| Theorem | funco 6606 | The composition of two functions is a function. Exercise 29 of [TakeutiZaring] p. 25. (Contributed by NM, 26-Jan-1997.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
| ⊢ ((Fun 𝐹 ∧ Fun 𝐺) → Fun (𝐹 ∘ 𝐺)) | ||
| Theorem | funresfunco 6607 | Composition of two functions, generalization of funco 6606. (Contributed by Alexander van der Vekens, 25-Jul-2017.) |
| ⊢ ((Fun (𝐹 ↾ ran 𝐺) ∧ Fun 𝐺) → Fun (𝐹 ∘ 𝐺)) | ||
| Theorem | funres 6608 | A restriction of a function is a function. Compare Exercise 18 of [TakeutiZaring] p. 25. (Contributed by NM, 16-Aug-1994.) |
| ⊢ (Fun 𝐹 → Fun (𝐹 ↾ 𝐴)) | ||
| Theorem | funresd 6609 | A restriction of a function is a function. (Contributed by Glauco Siliprandi, 2-Jan-2022.) |
| ⊢ (𝜑 → Fun 𝐹) ⇒ ⊢ (𝜑 → Fun (𝐹 ↾ 𝐴)) | ||
| Theorem | funssres 6610 | The restriction of a function to the domain of a subclass equals the subclass. (Contributed by NM, 15-Aug-1994.) |
| ⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹) → (𝐹 ↾ dom 𝐺) = 𝐺) | ||
| Theorem | fun2ssres 6611 | Equality of restrictions of a function and a subclass. (Contributed by NM, 16-Aug-1994.) |
| ⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹 ∧ 𝐴 ⊆ dom 𝐺) → (𝐹 ↾ 𝐴) = (𝐺 ↾ 𝐴)) | ||
| Theorem | funun 6612 | The union of functions with disjoint domains is a function. Theorem 4.6 of [Monk1] p. 43. (Contributed by NM, 12-Aug-1994.) |
| ⊢ (((Fun 𝐹 ∧ Fun 𝐺) ∧ (dom 𝐹 ∩ dom 𝐺) = ∅) → Fun (𝐹 ∪ 𝐺)) | ||
| Theorem | fununmo 6613* | If the union of classes is a function, there is at most one element in relation to an arbitrary element regarding one of these classes. (Contributed by AV, 18-Jul-2019.) |
| ⊢ (Fun (𝐹 ∪ 𝐺) → ∃*𝑦 𝑥𝐹𝑦) | ||
| Theorem | fununfun 6614 | If the union of classes is a function, the classes itselves are functions. (Contributed by AV, 18-Jul-2019.) |
| ⊢ (Fun (𝐹 ∪ 𝐺) → (Fun 𝐹 ∧ Fun 𝐺)) | ||
| Theorem | fundif 6615 | A function with removed elements is still a function. (Contributed by AV, 7-Jun-2021.) |
| ⊢ (Fun 𝐹 → Fun (𝐹 ∖ 𝐴)) | ||
| Theorem | funcnvsn 6616 | The converse singleton of an ordered pair is a function. This is equivalent to funsn 6619 via cnvsn 6246, but stating it this way allows to skip the sethood assumptions on 𝐴 and 𝐵. (Contributed by NM, 30-Apr-2015.) |
| ⊢ Fun ◡{〈𝐴, 𝐵〉} | ||
| Theorem | funsng 6617 | A singleton of an ordered pair is a function. Theorem 10.5 of [Quine] p. 65. (Contributed by NM, 28-Jun-2011.) |
| ⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → Fun {〈𝐴, 𝐵〉}) | ||
| Theorem | fnsng 6618 | Functionality and domain of the singleton of an ordered pair. (Contributed by Mario Carneiro, 30-Apr-2015.) |
| ⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → {〈𝐴, 𝐵〉} Fn {𝐴}) | ||
| Theorem | funsn 6619 | A singleton of an ordered pair is a function. Theorem 10.5 of [Quine] p. 65. (Contributed by NM, 12-Aug-1994.) |
| ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ Fun {〈𝐴, 𝐵〉} | ||
| Theorem | funprg 6620 | A set of two pairs is a function if their first members are different. (Contributed by FL, 26-Jun-2011.) (Proof shortened by JJ, 14-Jul-2021.) |
| ⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) | ||
| Theorem | funtpg 6621 | A set of three pairs is a function if their first members are different. (Contributed by Alexander van der Vekens, 5-Dec-2017.) (Proof shortened by JJ, 14-Jul-2021.) |
| ⊢ (((𝑋 ∈ 𝑈 ∧ 𝑌 ∈ 𝑉 ∧ 𝑍 ∈ 𝑊) ∧ (𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐺 ∧ 𝐶 ∈ 𝐻) ∧ (𝑋 ≠ 𝑌 ∧ 𝑋 ≠ 𝑍 ∧ 𝑌 ≠ 𝑍)) → Fun {〈𝑋, 𝐴〉, 〈𝑌, 𝐵〉, 〈𝑍, 𝐶〉}) | ||
| Theorem | funpr 6622 | A function with a domain of two elements. (Contributed by Jeff Madsen, 20-Jun-2010.) |
| ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V ⇒ ⊢ (𝐴 ≠ 𝐵 → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) | ||
| Theorem | funtp 6623 | A function with a domain of three elements. (Contributed by NM, 14-Sep-2011.) |
| ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V & ⊢ 𝐸 ∈ V & ⊢ 𝐹 ∈ V ⇒ ⊢ ((𝐴 ≠ 𝐵 ∧ 𝐴 ≠ 𝐶 ∧ 𝐵 ≠ 𝐶) → Fun {〈𝐴, 𝐷〉, 〈𝐵, 𝐸〉, 〈𝐶, 𝐹〉}) | ||
| Theorem | fnsn 6624 | Functionality and domain of the singleton of an ordered pair. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
| ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ {〈𝐴, 𝐵〉} Fn {𝐴} | ||
| Theorem | fnprg 6625 | Function with a domain of two different values. (Contributed by FL, 26-Jun-2011.) (Revised by Mario Carneiro, 26-Apr-2015.) |
| ⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉} Fn {𝐴, 𝐵}) | ||
| Theorem | fntpg 6626 | Function with a domain of three different values. (Contributed by Alexander van der Vekens, 5-Dec-2017.) |
| ⊢ (((𝑋 ∈ 𝑈 ∧ 𝑌 ∈ 𝑉 ∧ 𝑍 ∈ 𝑊) ∧ (𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐺 ∧ 𝐶 ∈ 𝐻) ∧ (𝑋 ≠ 𝑌 ∧ 𝑋 ≠ 𝑍 ∧ 𝑌 ≠ 𝑍)) → {〈𝑋, 𝐴〉, 〈𝑌, 𝐵〉, 〈𝑍, 𝐶〉} Fn {𝑋, 𝑌, 𝑍}) | ||
| Theorem | fntp 6627 | A function with a domain of three elements. (Contributed by NM, 14-Sep-2011.) (Revised by Mario Carneiro, 26-Apr-2015.) |
| ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V & ⊢ 𝐸 ∈ V & ⊢ 𝐹 ∈ V ⇒ ⊢ ((𝐴 ≠ 𝐵 ∧ 𝐴 ≠ 𝐶 ∧ 𝐵 ≠ 𝐶) → {〈𝐴, 𝐷〉, 〈𝐵, 𝐸〉, 〈𝐶, 𝐹〉} Fn {𝐴, 𝐵, 𝐶}) | ||
| Theorem | funcnvpr 6628 | The converse pair of ordered pairs is a function if the second members are different. Note that the second members need not be sets. (Contributed by AV, 23-Jan-2021.) |
| ⊢ ((𝐴 ∈ 𝑈 ∧ 𝐶 ∈ 𝑉 ∧ 𝐵 ≠ 𝐷) → Fun ◡{〈𝐴, 𝐵〉, 〈𝐶, 𝐷〉}) | ||
| Theorem | funcnvtp 6629 | The converse triple of ordered pairs is a function if the second members are pairwise different. Note that the second members need not be sets. (Contributed by AV, 23-Jan-2021.) |
| ⊢ (((𝐴 ∈ 𝑈 ∧ 𝐶 ∈ 𝑉 ∧ 𝐸 ∈ 𝑊) ∧ (𝐵 ≠ 𝐷 ∧ 𝐵 ≠ 𝐹 ∧ 𝐷 ≠ 𝐹)) → Fun ◡{〈𝐴, 𝐵〉, 〈𝐶, 𝐷〉, 〈𝐸, 𝐹〉}) | ||
| Theorem | funcnvqp 6630 | The converse quadruple of ordered pairs is a function if the second members are pairwise different. Note that the second members need not be sets. (Contributed by AV, 23-Jan-2021.) (Proof shortened by JJ, 14-Jul-2021.) |
| ⊢ ((((𝐴 ∈ 𝑈 ∧ 𝐶 ∈ 𝑉) ∧ (𝐸 ∈ 𝑊 ∧ 𝐺 ∈ 𝑇)) ∧ ((𝐵 ≠ 𝐷 ∧ 𝐵 ≠ 𝐹 ∧ 𝐵 ≠ 𝐻) ∧ (𝐷 ≠ 𝐹 ∧ 𝐷 ≠ 𝐻) ∧ 𝐹 ≠ 𝐻)) → Fun ◡({〈𝐴, 𝐵〉, 〈𝐶, 𝐷〉} ∪ {〈𝐸, 𝐹〉, 〈𝐺, 𝐻〉})) | ||
| Theorem | fun0 6631 | The empty set is a function. Theorem 10.3 of [Quine] p. 65. (Contributed by NM, 7-Apr-1998.) |
| ⊢ Fun ∅ | ||
| Theorem | funcnv0 6632 | The converse of the empty set is a function. (Contributed by AV, 7-Jan-2021.) |
| ⊢ Fun ◡∅ | ||
| Theorem | funcnvcnv 6633 | The double converse of a function is a function. (Contributed by NM, 21-Sep-2004.) |
| ⊢ (Fun 𝐴 → Fun ◡◡𝐴) | ||
| Theorem | funcnv2 6634* | A simpler equivalence for single-rooted (see funcnv 6635). (Contributed by NM, 9-Aug-2004.) |
| ⊢ (Fun ◡𝐴 ↔ ∀𝑦∃*𝑥 𝑥𝐴𝑦) | ||
| Theorem | funcnv 6635* | The converse of a class is a function iff the class is single-rooted, which means that for any 𝑦 in the range of 𝐴 there is at most one 𝑥 such that 𝑥𝐴𝑦. Definition of single-rooted in [Enderton] p. 43. See funcnv2 6634 for a simpler version. (Contributed by NM, 13-Aug-2004.) |
| ⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃*𝑥 𝑥𝐴𝑦) | ||
| Theorem | funcnv3 6636* | A condition showing a class is single-rooted. (See funcnv 6635). (Contributed by NM, 26-May-2006.) |
| ⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃!𝑥 ∈ dom 𝐴 𝑥𝐴𝑦) | ||
| Theorem | fun2cnv 6637* | The double converse of a class is a function iff the class is single-valued. Each side is equivalent to Definition 6.4(2) of [TakeutiZaring] p. 23, who use the notation "Un(A)" for single-valued. Note that 𝐴 is not necessarily a function. (Contributed by NM, 13-Aug-2004.) |
| ⊢ (Fun ◡◡𝐴 ↔ ∀𝑥∃*𝑦 𝑥𝐴𝑦) | ||
| Theorem | svrelfun 6638 | A single-valued relation is a function. (See fun2cnv 6637 for "single-valued.") Definition 6.4(4) of [TakeutiZaring] p. 24. (Contributed by NM, 17-Jan-2006.) |
| ⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ Fun ◡◡𝐴)) | ||
| Theorem | fncnv 6639* | Single-rootedness (see funcnv 6635) of a class cut down by a Cartesian product. (Contributed by NM, 5-Mar-2007.) |
| ⊢ (◡(𝑅 ∩ (𝐴 × 𝐵)) Fn 𝐵 ↔ ∀𝑦 ∈ 𝐵 ∃!𝑥 ∈ 𝐴 𝑥𝑅𝑦) | ||
| Theorem | fun11 6640* | Two ways of stating that 𝐴 is one-to-one (but not necessarily a function). Each side is equivalent to Definition 6.4(3) of [TakeutiZaring] p. 24, who use the notation "Un2 (A)" for one-to-one (but not necessarily a function). (Contributed by NM, 17-Jan-2006.) |
| ⊢ ((Fun ◡◡𝐴 ∧ Fun ◡𝐴) ↔ ∀𝑥∀𝑦∀𝑧∀𝑤((𝑥𝐴𝑦 ∧ 𝑧𝐴𝑤) → (𝑥 = 𝑧 ↔ 𝑦 = 𝑤))) | ||
| Theorem | fununi 6641* | The union of a chain (with respect to inclusion) of functions is a function. (Contributed by NM, 10-Aug-2004.) |
| ⊢ (∀𝑓 ∈ 𝐴 (Fun 𝑓 ∧ ∀𝑔 ∈ 𝐴 (𝑓 ⊆ 𝑔 ∨ 𝑔 ⊆ 𝑓)) → Fun ∪ 𝐴) | ||
| Theorem | funin 6642 | The intersection with a function is a function. Exercise 14(a) of [Enderton] p. 53. (Contributed by NM, 19-Mar-2004.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
| ⊢ (Fun 𝐹 → Fun (𝐹 ∩ 𝐺)) | ||
| Theorem | funres11 6643 | The restriction of a one-to-one function is one-to-one. (Contributed by NM, 25-Mar-1998.) |
| ⊢ (Fun ◡𝐹 → Fun ◡(𝐹 ↾ 𝐴)) | ||
| Theorem | funcnvres 6644 | The converse of a restricted function. (Contributed by NM, 27-Mar-1998.) |
| ⊢ (Fun ◡𝐹 → ◡(𝐹 ↾ 𝐴) = (◡𝐹 ↾ (𝐹 “ 𝐴))) | ||
| Theorem | cnvresid 6645 | Converse of a restricted identity function. (Contributed by FL, 4-Mar-2007.) |
| ⊢ ◡( I ↾ 𝐴) = ( I ↾ 𝐴) | ||
| Theorem | funcnvres2 6646 | The converse of a restriction of the converse of a function equals the function restricted to the image of its converse. (Contributed by NM, 4-May-2005.) |
| ⊢ (Fun 𝐹 → ◡(◡𝐹 ↾ 𝐴) = (𝐹 ↾ (◡𝐹 “ 𝐴))) | ||
| Theorem | funimacnv 6647 | The image of the preimage of a function. (Contributed by NM, 25-May-2004.) |
| ⊢ (Fun 𝐹 → (𝐹 “ (◡𝐹 “ 𝐴)) = (𝐴 ∩ ran 𝐹)) | ||
| Theorem | funimass1 6648 | A kind of contraposition law that infers a subclass of an image from a preimage subclass. (Contributed by NM, 25-May-2004.) |
| ⊢ ((Fun 𝐹 ∧ 𝐴 ⊆ ran 𝐹) → ((◡𝐹 “ 𝐴) ⊆ 𝐵 → 𝐴 ⊆ (𝐹 “ 𝐵))) | ||
| Theorem | funimass2 6649 | A kind of contraposition law that infers an image subclass from a subclass of a preimage. (Contributed by NM, 25-May-2004.) |
| ⊢ ((Fun 𝐹 ∧ 𝐴 ⊆ (◡𝐹 “ 𝐵)) → (𝐹 “ 𝐴) ⊆ 𝐵) | ||
| Theorem | imadif 6650 | The image of a difference is the difference of images. (Contributed by NM, 24-May-1998.) |
| ⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∖ 𝐵)) = ((𝐹 “ 𝐴) ∖ (𝐹 “ 𝐵))) | ||
| Theorem | imain 6651 | The image of an intersection is the intersection of images. (Contributed by Paul Chapman, 11-Apr-2009.) |
| ⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∩ 𝐵)) = ((𝐹 “ 𝐴) ∩ (𝐹 “ 𝐵))) | ||
| Theorem | f1imadifssran 6652 | Condition for the range of a one-to-one function to be the range of one its restrictions. Variant of imadifssran 6171. (Contributed by AV, 4-Oct-2025.) |
| ⊢ (Fun ◡𝐹 → ((𝐹 “ (dom 𝐹 ∖ 𝐴)) ⊆ ran (𝐹 ↾ 𝐴) → ran 𝐹 = ran (𝐹 ↾ 𝐴))) | ||
| Theorem | funimaexg 6653 | Axiom of Replacement using abbreviations. Axiom 39(vi) of [Quine] p. 284. Compare Exercise 9 of [TakeutiZaring] p. 29. (Contributed by NM, 10-Sep-2006.) Shorten proof and avoid ax-10 2141, ax-12 2177. (Revised by SN, 19-Dec-2024.) |
| ⊢ ((Fun 𝐴 ∧ 𝐵 ∈ 𝐶) → (𝐴 “ 𝐵) ∈ V) | ||
| Theorem | funimaexgOLD 6654 | Obsolete version of funimaexg 6653 as of 19-Dec-2024. (Contributed by NM, 10-Sep-2006.) (Proof modification is discouraged.) (New usage is discouraged.) |
| ⊢ ((Fun 𝐴 ∧ 𝐵 ∈ 𝐶) → (𝐴 “ 𝐵) ∈ V) | ||
| Theorem | funimaex 6655 | The image of a set under any function is also a set. Equivalent of Axiom of Replacement ax-rep 5279. 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 6656* | 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.) (Proof shortened by SN, 19-Dec-2024.) |
| ⊢ (𝑏 ∈ ({〈𝑥, 𝑦〉 ∣ 𝜑} “ 𝐴) ↔ ∃𝑥 ∈ 𝐴 [𝑏 / 𝑦]𝜑) | ||
| Theorem | isarep1OLD 6657* | Obsolete version of isarep1 6656 as of 19-Dec-2024. (Contributed by NM, 26-Oct-2006.) (Proof shortened by Mario Carneiro, 4-Dec-2016.) (Proof modification is discouraged.) (New usage is discouraged.) |
| ⊢ (𝑏 ∈ ({〈𝑥, 𝑦〉 ∣ 𝜑} “ 𝐴) ↔ ∃𝑥 ∈ 𝐴 [𝑏 / 𝑦]𝜑) | ||
| Theorem | isarep2 6658* | 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 6655. (Contributed by NM, 26-Oct-2006.) |
| ⊢ 𝐴 ∈ V & ⊢ ∀𝑥 ∈ 𝐴 ∀𝑦∀𝑧((𝜑 ∧ [𝑧 / 𝑦]𝜑) → 𝑦 = 𝑧) ⇒ ⊢ ∃𝑤 𝑤 = ({〈𝑥, 𝑦〉 ∣ 𝜑} “ 𝐴) | ||
| Theorem | fneq1 6659 | Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.) |
| ⊢ (𝐹 = 𝐺 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) | ||
| Theorem | fneq2 6660 | Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.) |
| ⊢ (𝐴 = 𝐵 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) | ||
| Theorem | fneq1d 6661 | Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
| ⊢ (𝜑 → 𝐹 = 𝐺) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) | ||
| Theorem | fneq2d 6662 | Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
| ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) | ||
| Theorem | fneq12d 6663 | Equality deduction for function predicate with domain. (Contributed by NM, 26-Jun-2011.) |
| ⊢ (𝜑 → 𝐹 = 𝐺) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) | ||
| Theorem | fneq12 6664 | Equality theorem for function predicate with domain. (Contributed by Thierry Arnoux, 31-Jan-2017.) |
| ⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐵) → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) | ||
| Theorem | fneq1i 6665 | Equality inference for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
| ⊢ 𝐹 = 𝐺 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴) | ||
| Theorem | fneq2i 6666 | Equality inference for function predicate with domain. (Contributed by NM, 4-Sep-2011.) |
| ⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵) | ||
| Theorem | nffn 6667 | Bound-variable hypothesis builder for a function with domain. (Contributed by NM, 30-Jan-2004.) |
| ⊢ Ⅎ𝑥𝐹 & ⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥 𝐹 Fn 𝐴 | ||
| Theorem | fnfun 6668 | A function with domain is a function. (Contributed by NM, 1-Aug-1994.) |
| ⊢ (𝐹 Fn 𝐴 → Fun 𝐹) | ||
| Theorem | fnfund 6669 | A function with domain is a function, deduction form. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
| ⊢ (𝜑 → 𝐹 Fn 𝐴) ⇒ ⊢ (𝜑 → Fun 𝐹) | ||
| Theorem | fnrel 6670 | A function with domain is a relation. (Contributed by NM, 1-Aug-1994.) |
| ⊢ (𝐹 Fn 𝐴 → Rel 𝐹) | ||
| Theorem | fndm 6671 | The domain of a function. (Contributed by NM, 2-Aug-1994.) |
| ⊢ (𝐹 Fn 𝐴 → dom 𝐹 = 𝐴) | ||
| Theorem | fndmi 6672 | The domain of a function. (Contributed by Wolf Lammen, 1-Jun-2024.) |
| ⊢ 𝐹 Fn 𝐴 ⇒ ⊢ dom 𝐹 = 𝐴 | ||
| Theorem | fndmd 6673 | The domain of a function. (Contributed by Glauco Siliprandi, 23-Oct-2021.) |
| ⊢ (𝜑 → 𝐹 Fn 𝐴) ⇒ ⊢ (𝜑 → dom 𝐹 = 𝐴) | ||
| Theorem | funfni 6674 | Inference to convert a function and domain antecedent. (Contributed by NM, 22-Apr-2004.) |
| ⊢ ((Fun 𝐹 ∧ 𝐵 ∈ dom 𝐹) → 𝜑) ⇒ ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → 𝜑) | ||
| Theorem | fndmu 6675 | A function has a unique domain. (Contributed by NM, 11-Aug-1994.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐹 Fn 𝐵) → 𝐴 = 𝐵) | ||
| Theorem | fnbr 6676 | The first argument of binary relation on a function belongs to the function's domain. (Contributed by NM, 7-May-2004.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵𝐹𝐶) → 𝐵 ∈ 𝐴) | ||
| Theorem | fnop 6677 | 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 6678* | 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 6679* | There is exactly one value of a function. (Contributed by NM, 7-Nov-1995.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → ∃!𝑦〈𝐵, 𝑦〉 ∈ 𝐹) | ||
| Theorem | fnunres1 6680 | Restriction of a disjoint union to the domain of the first function. (Contributed by Thierry Arnoux, 9-Dec-2021.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵 ∧ (𝐴 ∩ 𝐵) = ∅) → ((𝐹 ∪ 𝐺) ↾ 𝐴) = 𝐹) | ||
| Theorem | fnunres2 6681 | Restriction of a disjoint union to the domain of the second function. (Contributed by Thierry Arnoux, 12-Oct-2023.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵 ∧ (𝐴 ∩ 𝐵) = ∅) → ((𝐹 ∪ 𝐺) ↾ 𝐵) = 𝐺) | ||
| Theorem | fnun 6682 | The union of two functions with disjoint domains. (Contributed by NM, 22-Sep-2004.) |
| ⊢ (((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐹 ∪ 𝐺) Fn (𝐴 ∪ 𝐵)) | ||
| Theorem | fnund 6683 | The union of two functions with disjoint domains, a deduction version. (Contributed by metakunt, 28-May-2024.) |
| ⊢ (𝜑 → 𝐹 Fn 𝐴) & ⊢ (𝜑 → 𝐺 Fn 𝐵) & ⊢ (𝜑 → (𝐴 ∩ 𝐵) = ∅) ⇒ ⊢ (𝜑 → (𝐹 ∪ 𝐺) Fn (𝐴 ∪ 𝐵)) | ||
| Theorem | fnunop 6684 | Extension of a function with a new ordered pair. (Contributed by NM, 28-Sep-2013.) (Revised by Mario Carneiro, 30-Apr-2015.) (Revised by AV, 16-Aug-2024.) |
| ⊢ (𝜑 → 𝑋 ∈ 𝑉) & ⊢ (𝜑 → 𝑌 ∈ 𝑊) & ⊢ (𝜑 → 𝐹 Fn 𝐷) & ⊢ 𝐺 = (𝐹 ∪ {〈𝑋, 𝑌〉}) & ⊢ 𝐸 = (𝐷 ∪ {𝑋}) & ⊢ (𝜑 → ¬ 𝑋 ∈ 𝐷) ⇒ ⊢ (𝜑 → 𝐺 Fn 𝐸) | ||
| Theorem | fncofn 6685 | Composition of a function with domain and a function as a function with domain. Generalization of fnco 6686. (Contributed by AV, 17-Sep-2024.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ Fun 𝐺) → (𝐹 ∘ 𝐺) Fn (◡𝐺 “ 𝐴)) | ||
| Theorem | fnco 6686 | Composition of two functions with domains as a function with domain. (Contributed by NM, 22-May-2006.) (Proof shortened by AV, 20-Sep-2024.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵 ∧ ran 𝐺 ⊆ 𝐴) → (𝐹 ∘ 𝐺) Fn 𝐵) | ||
| Theorem | fnresdm 6687 | A function does not change when restricted to its domain. (Contributed by NM, 5-Sep-2004.) |
| ⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ 𝐴) = 𝐹) | ||
| Theorem | fnresdisj 6688 | A function restricted to a class disjoint with its domain is empty. (Contributed by NM, 23-Sep-2004.) |
| ⊢ (𝐹 Fn 𝐴 → ((𝐴 ∩ 𝐵) = ∅ ↔ (𝐹 ↾ 𝐵) = ∅)) | ||
| Theorem | 2elresin 6689 | Membership in two functions restricted by each other's domain. (Contributed by NM, 8-Aug-1994.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) → ((〈𝑥, 𝑦〉 ∈ 𝐹 ∧ 〈𝑥, 𝑧〉 ∈ 𝐺) ↔ (〈𝑥, 𝑦〉 ∈ (𝐹 ↾ (𝐴 ∩ 𝐵)) ∧ 〈𝑥, 𝑧〉 ∈ (𝐺 ↾ (𝐴 ∩ 𝐵))))) | ||
| Theorem | fnssresb 6690 | Restriction of a function with a subclass of its domain. (Contributed by NM, 10-Oct-2007.) |
| ⊢ (𝐹 Fn 𝐴 → ((𝐹 ↾ 𝐵) Fn 𝐵 ↔ 𝐵 ⊆ 𝐴)) | ||
| Theorem | fnssres 6691 | Restriction of a function with a subclass of its domain. (Contributed by NM, 2-Aug-1994.) |
| ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ⊆ 𝐴) → (𝐹 ↾ 𝐵) Fn 𝐵) | ||
| Theorem | fnssresd 6692 | Restriction of a function to a subclass of its domain. (Contributed by Glauco Siliprandi, 5-Feb-2022.) |
| ⊢ (𝜑 → 𝐹 Fn 𝐴) & ⊢ (𝜑 → 𝐵 ⊆ 𝐴) ⇒ ⊢ (𝜑 → (𝐹 ↾ 𝐵) Fn 𝐵) | ||
| Theorem | fnresin1 6693 | Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.) |
| ⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐴 ∩ 𝐵)) Fn (𝐴 ∩ 𝐵)) | ||
| Theorem | fnresin2 6694 | Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.) |
| ⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐵 ∩ 𝐴)) Fn (𝐵 ∩ 𝐴)) | ||
| Theorem | fnres 6695* | An equivalence for functionality of a restriction. Compare dffun8 6594. (Contributed by Mario Carneiro, 20-May-2015.) (Proof shortened by Peter Mazsa, 2-Oct-2022.) |
| ⊢ ((𝐹 ↾ 𝐴) Fn 𝐴 ↔ ∀𝑥 ∈ 𝐴 ∃!𝑦 𝑥𝐹𝑦) | ||
| Theorem | idfn 6696 | The identity relation is a function on the universal class. See also funi 6598. (Contributed by BJ, 23-Dec-2023.) |
| ⊢ I Fn V | ||
| Theorem | fnresi 6697 | The restricted identity relation is a function on the restricting class. (Contributed by NM, 27-Aug-2004.) (Proof shortened by BJ, 27-Dec-2023.) |
| ⊢ ( I ↾ 𝐴) Fn 𝐴 | ||
| Theorem | fnima 6698 | 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 6699 | A function with empty domain is empty. (Contributed by NM, 15-Apr-1998.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
| ⊢ (𝐹 Fn ∅ ↔ 𝐹 = ∅) | ||
| Theorem | fnimadisj 6700 | 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 𝐴 ∧ (𝐴 ∩ 𝐶) = ∅) → (𝐹 “ 𝐶) = ∅) | ||
| < Previous Next > |
| Copyright terms: Public domain | < Previous Next > |