![]() |
Metamath
Proof Explorer Theorem List (p. 61 of 429) | < 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-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | fun11 6001* | 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 6002* | The union of a chain (with respect to inclusion) of functions is a function. (Contributed by NM, 10-Aug-2004.) |
⊢ (∀𝑓 ∈ 𝐴 (Fun 𝑓 ∧ ∀𝑔 ∈ 𝐴 (𝑓 ⊆ 𝑔 ∨ 𝑔 ⊆ 𝑓)) → Fun ∪ 𝐴) | ||
Theorem | funin 6003 | 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 6004 | The restriction of a one-to-one function is one-to-one. (Contributed by NM, 25-Mar-1998.) |
⊢ (Fun ◡𝐹 → Fun ◡(𝐹 ↾ 𝐴)) | ||
Theorem | funcnvres 6005 | The converse of a restricted function. (Contributed by NM, 27-Mar-1998.) |
⊢ (Fun ◡𝐹 → ◡(𝐹 ↾ 𝐴) = (◡𝐹 ↾ (𝐹 “ 𝐴))) | ||
Theorem | cnvresid 6006 | Converse of a restricted identity function. (Contributed by FL, 4-Mar-2007.) |
⊢ ◡( I ↾ 𝐴) = ( I ↾ 𝐴) | ||
Theorem | funcnvres2 6007 | 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 6008 | The image of the preimage of a function. (Contributed by NM, 25-May-2004.) |
⊢ (Fun 𝐹 → (𝐹 “ (◡𝐹 “ 𝐴)) = (𝐴 ∩ ran 𝐹)) | ||
Theorem | funimass1 6009 | 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 6010 | 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 6011 | The image of a difference is the difference of images. (Contributed by NM, 24-May-1998.) |
⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∖ 𝐵)) = ((𝐹 “ 𝐴) ∖ (𝐹 “ 𝐵))) | ||
Theorem | imain 6012 | The image of an intersection is the intersection of images. (Contributed by Paul Chapman, 11-Apr-2009.) |
⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∩ 𝐵)) = ((𝐹 “ 𝐴) ∩ (𝐹 “ 𝐵))) | ||
Theorem | funimaexg 6013 | 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.) |
⊢ ((Fun 𝐴 ∧ 𝐵 ∈ 𝐶) → (𝐴 “ 𝐵) ∈ V) | ||
Theorem | funimaex 6014 | The image of a set under any function is also a set. Equivalent of Axiom of Replacement ax-rep 4804. 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 6015* | 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 6016* | 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 6014. (Contributed by NM, 26-Oct-2006.) |
⊢ 𝐴 ∈ V & ⊢ ∀𝑥 ∈ 𝐴 ∀𝑦∀𝑧((𝜑 ∧ [𝑧 / 𝑦]𝜑) → 𝑦 = 𝑧) ⇒ ⊢ ∃𝑤 𝑤 = ({〈𝑥, 𝑦〉 ∣ 𝜑} “ 𝐴) | ||
Theorem | fneq1 6017 | Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹 = 𝐺 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) | ||
Theorem | fneq2 6018 | Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐴 = 𝐵 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) | ||
Theorem | fneq1d 6019 | Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → 𝐹 = 𝐺) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) | ||
Theorem | fneq2d 6020 | Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) | ||
Theorem | fneq12d 6021 | Equality deduction for function predicate with domain. (Contributed by NM, 26-Jun-2011.) |
⊢ (𝜑 → 𝐹 = 𝐺) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) | ||
Theorem | fneq12 6022 | Equality theorem for function predicate with domain. (Contributed by Thierry Arnoux, 31-Jan-2017.) |
⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐵) → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) | ||
Theorem | fneq1i 6023 | Equality inference for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ 𝐹 = 𝐺 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴) | ||
Theorem | fneq2i 6024 | Equality inference for function predicate with domain. (Contributed by NM, 4-Sep-2011.) |
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵) | ||
Theorem | nffn 6025 | Bound-variable hypothesis builder for a function with domain. (Contributed by NM, 30-Jan-2004.) |
⊢ Ⅎ𝑥𝐹 & ⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥 𝐹 Fn 𝐴 | ||
Theorem | fnfun 6026 | A function with domain is a function. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹 Fn 𝐴 → Fun 𝐹) | ||
Theorem | fnrel 6027 | A function with domain is a relation. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹 Fn 𝐴 → Rel 𝐹) | ||
Theorem | fndm 6028 | The domain of a function. (Contributed by NM, 2-Aug-1994.) |
⊢ (𝐹 Fn 𝐴 → dom 𝐹 = 𝐴) | ||
Theorem | funfni 6029 | Inference to convert a function and domain antecedent. (Contributed by NM, 22-Apr-2004.) |
⊢ ((Fun 𝐹 ∧ 𝐵 ∈ dom 𝐹) → 𝜑) ⇒ ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → 𝜑) | ||
Theorem | fndmu 6030 | A function has a unique domain. (Contributed by NM, 11-Aug-1994.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐹 Fn 𝐵) → 𝐴 = 𝐵) | ||
Theorem | fnbr 6031 | The first argument of binary relation on a function belongs to the function's domain. (Contributed by NM, 7-May-2004.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵𝐹𝐶) → 𝐵 ∈ 𝐴) | ||
Theorem | fnop 6032 | 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 6033* | 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 6034* | There is exactly one value of a function. (Contributed by NM, 7-Nov-1995.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → ∃!𝑦〈𝐵, 𝑦〉 ∈ 𝐹) | ||
Theorem | fnun 6035 | The union of two functions with disjoint domains. (Contributed by NM, 22-Sep-2004.) |
⊢ (((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐹 ∪ 𝐺) Fn (𝐴 ∪ 𝐵)) | ||
Theorem | fnunsn 6036 | 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 6037 | Composition of two functions. (Contributed by NM, 22-May-2006.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵 ∧ ran 𝐺 ⊆ 𝐴) → (𝐹 ∘ 𝐺) Fn 𝐵) | ||
Theorem | fnresdm 6038 | A function does not change when restricted to its domain. (Contributed by NM, 5-Sep-2004.) |
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ 𝐴) = 𝐹) | ||
Theorem | fnresdisj 6039 | A function restricted to a class disjoint with its domain is empty. (Contributed by NM, 23-Sep-2004.) |
⊢ (𝐹 Fn 𝐴 → ((𝐴 ∩ 𝐵) = ∅ ↔ (𝐹 ↾ 𝐵) = ∅)) | ||
Theorem | 2elresin 6040 | Membership in two functions restricted by each other's domain. (Contributed by NM, 8-Aug-1994.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐵) → ((〈𝑥, 𝑦〉 ∈ 𝐹 ∧ 〈𝑥, 𝑧〉 ∈ 𝐺) ↔ (〈𝑥, 𝑦〉 ∈ (𝐹 ↾ (𝐴 ∩ 𝐵)) ∧ 〈𝑥, 𝑧〉 ∈ (𝐺 ↾ (𝐴 ∩ 𝐵))))) | ||
Theorem | fnssresb 6041 | Restriction of a function with a subclass of its domain. (Contributed by NM, 10-Oct-2007.) |
⊢ (𝐹 Fn 𝐴 → ((𝐹 ↾ 𝐵) Fn 𝐵 ↔ 𝐵 ⊆ 𝐴)) | ||
Theorem | fnssres 6042 | Restriction of a function with a subclass of its domain. (Contributed by NM, 2-Aug-1994.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ⊆ 𝐴) → (𝐹 ↾ 𝐵) Fn 𝐵) | ||
Theorem | fnresin1 6043 | Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.) |
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐴 ∩ 𝐵)) Fn (𝐴 ∩ 𝐵)) | ||
Theorem | fnresin2 6044 | Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.) |
⊢ (𝐹 Fn 𝐴 → (𝐹 ↾ (𝐵 ∩ 𝐴)) Fn (𝐵 ∩ 𝐴)) | ||
Theorem | fnres 6045* | An equivalence for functionality of a restriction. Compare dffun8 5954. (Contributed by Mario Carneiro, 20-May-2015.) |
⊢ ((𝐹 ↾ 𝐴) Fn 𝐴 ↔ ∀𝑥 ∈ 𝐴 ∃!𝑦 𝑥𝐹𝑦) | ||
Theorem | fnresi 6046 | Functionality and domain of restricted identity. (Contributed by NM, 27-Aug-2004.) |
⊢ ( I ↾ 𝐴) Fn 𝐴 | ||
Theorem | idssxp 6047 | A diagonal set as a subset of a Cartesian product. (Contributed by Thierry Arnoux, 29-Dec-2019.) |
⊢ ( I ↾ 𝐴) ⊆ (𝐴 × 𝐴) | ||
Theorem | fnima 6048 | 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 6049 | A function with empty domain is empty. (Contributed by NM, 15-Apr-1998.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ (𝐹 Fn ∅ ↔ 𝐹 = ∅) | ||
Theorem | fnimadisj 6050 | 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 6051 | Images under a function never map nonempty sets to empty sets. EDITORIAL: usable in fnwe2lem2 37938. (Contributed by Stefan O'Rear, 21-Jan-2015.) |
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ⊆ 𝐴) → ((𝐹 “ 𝐵) = ∅ ↔ 𝐵 = ∅)) | ||
Theorem | dfmpt3 6052 | Alternate definition for the "maps to" notation df-mpt 4763. (Contributed by Mario Carneiro, 30-Dec-2016.) |
⊢ (𝑥 ∈ 𝐴 ↦ 𝐵) = ∪ 𝑥 ∈ 𝐴 ({𝑥} × {𝐵}) | ||
Theorem | mptfnf 6053 | The maps-to notation defines a function with domain. (Contributed by Scott Fenton, 21-Mar-2011.) (Revised by Thierry Arnoux, 10-May-2017.) |
⊢ Ⅎ𝑥𝐴 ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ V ↔ (𝑥 ∈ 𝐴 ↦ 𝐵) Fn 𝐴) | ||
Theorem | fnmptf 6054 | The maps-to notation defines a function with domain. (Contributed by NM, 9-Apr-2013.) (Revised by Thierry Arnoux, 10-May-2017.) |
⊢ Ⅎ𝑥𝐴 ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ 𝑉 → (𝑥 ∈ 𝐴 ↦ 𝐵) Fn 𝐴) | ||
Theorem | fnopabg 6055* | 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 6056* | Functionality and domain of an ordered-pair class abstraction. (Contributed by NM, 5-Mar-1996.) |
⊢ (𝑥 ∈ 𝐴 → ∃!𝑦𝜑) & ⊢ 𝐹 = {〈𝑥, 𝑦〉 ∣ (𝑥 ∈ 𝐴 ∧ 𝜑)} ⇒ ⊢ 𝐹 Fn 𝐴 | ||
Theorem | mptfng 6057* | The maps-to notation defines a function with domain. (Contributed by Scott Fenton, 21-Mar-2011.) |
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ V ↔ 𝐹 Fn 𝐴) | ||
Theorem | fnmpt 6058* | The maps-to notation defines a function with domain. (Contributed by NM, 9-Apr-2013.) |
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ 𝑉 → 𝐹 Fn 𝐴) | ||
Theorem | mpt0 6059 | A mapping operation with empty domain. (Contributed by Mario Carneiro, 28-Dec-2014.) |
⊢ (𝑥 ∈ ∅ ↦ 𝐴) = ∅ | ||
Theorem | fnmpti 6060* | 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 6061* | Domain of the mapping operation. (Contributed by NM, 6-Sep-2005.) (Revised by Mario Carneiro, 31-Aug-2015.) |
⊢ 𝐵 ∈ V & ⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) ⇒ ⊢ dom 𝐹 = 𝐴 | ||
Theorem | dmmptd 6062* | The domain of the mapping operation, deduction form. (Contributed by Glauco Siliprandi, 11-Dec-2019.) |
⊢ 𝐴 = (𝑥 ∈ 𝐵 ↦ 𝐶) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐵) → 𝐶 ∈ 𝑉) ⇒ ⊢ (𝜑 → dom 𝐴 = 𝐵) | ||
Theorem | mptun 6063 | Union of mappings which are mutually compatible. (Contributed by Mario Carneiro, 31-Aug-2015.) |
⊢ (𝑥 ∈ (𝐴 ∪ 𝐵) ↦ 𝐶) = ((𝑥 ∈ 𝐴 ↦ 𝐶) ∪ (𝑥 ∈ 𝐵 ↦ 𝐶)) | ||
Theorem | feq1 6064 | Equality theorem for functions. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹 = 𝐺 → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵)) | ||
Theorem | feq2 6065 | Equality theorem for functions. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐴 = 𝐵 → (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶)) | ||
Theorem | feq3 6066 | Equality theorem for functions. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐴 = 𝐵 → (𝐹:𝐶⟶𝐴 ↔ 𝐹:𝐶⟶𝐵)) | ||
Theorem | feq23 6067 | Equality theorem for functions. (Contributed by FL, 14-Jul-2007.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ ((𝐴 = 𝐶 ∧ 𝐵 = 𝐷) → (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷)) | ||
Theorem | feq1d 6068 | Equality deduction for functions. (Contributed by NM, 19-Feb-2008.) |
⊢ (𝜑 → 𝐹 = 𝐺) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵)) | ||
Theorem | feq2d 6069 | Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶)) | ||
Theorem | feq3d 6070 | Equality deduction for functions. (Contributed by AV, 1-Jan-2020.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝑋⟶𝐴 ↔ 𝐹:𝑋⟶𝐵)) | ||
Theorem | feq12d 6071 | Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → 𝐹 = 𝐺) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐺:𝐵⟶𝐶)) | ||
Theorem | feq123d 6072 | Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → 𝐹 = 𝐺) & ⊢ (𝜑 → 𝐴 = 𝐵) & ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐶 ↔ 𝐺:𝐵⟶𝐷)) | ||
Theorem | feq123 6073 | Equality theorem for functions. (Contributed by FL, 16-Nov-2008.) |
⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐶 ∧ 𝐵 = 𝐷) → (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐶⟶𝐷)) | ||
Theorem | feq1i 6074 | Equality inference for functions. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ 𝐹 = 𝐺 ⇒ ⊢ (𝐹:𝐴⟶𝐵 ↔ 𝐺:𝐴⟶𝐵) | ||
Theorem | feq2i 6075 | Equality inference for functions. (Contributed by NM, 5-Sep-2011.) |
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹:𝐴⟶𝐶 ↔ 𝐹:𝐵⟶𝐶) | ||
Theorem | feq12i 6076 | Equality inference for functions. (Contributed by AV, 7-Feb-2021.) |
⊢ 𝐹 = 𝐺 & ⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹:𝐴⟶𝐶 ↔ 𝐺:𝐵⟶𝐶) | ||
Theorem | feq23i 6077 | Equality inference for functions. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ 𝐴 = 𝐶 & ⊢ 𝐵 = 𝐷 ⇒ ⊢ (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷) | ||
Theorem | feq23d 6078 | Equality deduction for functions. (Contributed by NM, 8-Jun-2013.) |
⊢ (𝜑 → 𝐴 = 𝐶) & ⊢ (𝜑 → 𝐵 = 𝐷) ⇒ ⊢ (𝜑 → (𝐹:𝐴⟶𝐵 ↔ 𝐹:𝐶⟶𝐷)) | ||
Theorem | nff 6079 | Bound-variable hypothesis builder for a mapping. (Contributed by NM, 29-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.) |
⊢ Ⅎ𝑥𝐹 & ⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥 𝐹:𝐴⟶𝐵 | ||
Theorem | sbcfng 6080* | Distribute proper substitution through the function predicate with a domain. (Contributed by Alexander van der Vekens, 15-Jul-2018.) |
⊢ (𝑋 ∈ 𝑉 → ([𝑋 / 𝑥]𝐹 Fn 𝐴 ↔ ⦋𝑋 / 𝑥⦌𝐹 Fn ⦋𝑋 / 𝑥⦌𝐴)) | ||
Theorem | sbcfg 6081* | Distribute proper substitution through the function predicate with domain and codomain. (Contributed by Alexander van der Vekens, 15-Jul-2018.) |
⊢ (𝑋 ∈ 𝑉 → ([𝑋 / 𝑥]𝐹:𝐴⟶𝐵 ↔ ⦋𝑋 / 𝑥⦌𝐹:⦋𝑋 / 𝑥⦌𝐴⟶⦋𝑋 / 𝑥⦌𝐵)) | ||
Theorem | elimf 6082 | Eliminate a mapping hypothesis for the weak deduction theorem dedth 4172, when a special case 𝐺:𝐴⟶𝐵 is provable, in order to convert 𝐹:𝐴⟶𝐵 from a hypothesis to an antecedent. (Contributed by NM, 24-Aug-2006.) |
⊢ 𝐺:𝐴⟶𝐵 ⇒ ⊢ if(𝐹:𝐴⟶𝐵, 𝐹, 𝐺):𝐴⟶𝐵 | ||
Theorem | ffn 6083 | A mapping is a function with domain. (Contributed by NM, 2-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 → 𝐹 Fn 𝐴) | ||
Theorem | ffnd 6084 | A mapping is a function with domain, deduction form. (Contributed by Glauco Siliprandi, 17-Aug-2020.) |
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → 𝐹 Fn 𝐴) | ||
Theorem | dffn2 6085 | 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 6086 | A mapping is a function. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 → Fun 𝐹) | ||
Theorem | ffund 6087 | A mapping is a function, deduction version. (Contributed by Glauco Siliprandi, 3-Mar-2021.) |
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) ⇒ ⊢ (𝜑 → Fun 𝐹) | ||
Theorem | frel 6088 | A mapping is a relation. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 → Rel 𝐹) | ||
Theorem | fdm 6089 | The domain of a mapping. (Contributed by NM, 2-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 → dom 𝐹 = 𝐴) | ||
Theorem | fdmi 6090 | The domain of a mapping. (Contributed by NM, 28-Jul-2008.) |
⊢ 𝐹:𝐴⟶𝐵 ⇒ ⊢ dom 𝐹 = 𝐴 | ||
Theorem | frn 6091 | The range of a mapping. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 → ran 𝐹 ⊆ 𝐵) | ||
Theorem | dffn3 6092 | A function maps to its range. (Contributed by NM, 1-Sep-1999.) |
⊢ (𝐹 Fn 𝐴 ↔ 𝐹:𝐴⟶ran 𝐹) | ||
Theorem | ffrn 6093 | A function maps to its range. (Contributed by Glauco Siliprandi, 3-Mar-2021.) |
⊢ (𝐹:𝐴⟶𝐵 → 𝐹:𝐴⟶ran 𝐹) | ||
Theorem | fss 6094 | Expanding the codomain of a mapping. (Contributed by NM, 10-May-1998.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ ((𝐹:𝐴⟶𝐵 ∧ 𝐵 ⊆ 𝐶) → 𝐹:𝐴⟶𝐶) | ||
Theorem | fssd 6095 | Expanding the codomain of a mapping, deduction form. (Contributed by Glauco Siliprandi, 11-Dec-2019.) |
⊢ (𝜑 → 𝐹:𝐴⟶𝐵) & ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐹:𝐴⟶𝐶) | ||
Theorem | fco 6096 | Composition of two mappings. (Contributed by NM, 29-Aug-1999.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ ((𝐹:𝐵⟶𝐶 ∧ 𝐺:𝐴⟶𝐵) → (𝐹 ∘ 𝐺):𝐴⟶𝐶) | ||
Theorem | fco2 6097 | Functionality of a composition with weakened out of domain condition on the first argument. (Contributed by Stefan O'Rear, 11-Mar-2015.) |
⊢ (((𝐹 ↾ 𝐵):𝐵⟶𝐶 ∧ 𝐺:𝐴⟶𝐵) → (𝐹 ∘ 𝐺):𝐴⟶𝐶) | ||
Theorem | fssxp 6098 | A mapping is a class of ordered pairs. (Contributed by NM, 3-Aug-1994.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ (𝐹:𝐴⟶𝐵 → 𝐹 ⊆ (𝐴 × 𝐵)) | ||
Theorem | funssxp 6099 | Two ways of specifying a partial function from 𝐴 to 𝐵. (Contributed by NM, 13-Nov-2007.) |
⊢ ((Fun 𝐹 ∧ 𝐹 ⊆ (𝐴 × 𝐵)) ↔ (𝐹:dom 𝐹⟶𝐵 ∧ dom 𝐹 ⊆ 𝐴)) | ||
Theorem | ffdm 6100 | A mapping is a partial function. (Contributed by NM, 25-Nov-2007.) |
⊢ (𝐹:𝐴⟶𝐵 → (𝐹:dom 𝐹⟶𝐵 ∧ dom 𝐹 ⊆ 𝐴)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |