HomeHome Intuitionistic Logic Explorer
Theorem List (p. 54 of 149)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 5301-5400   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremfunimaex 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)
 
Theoremisarep1 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.)
(𝑏 ∈ ({⟨π‘₯, π‘¦βŸ© ∣ πœ‘} β€œ 𝐴) ↔ βˆƒπ‘₯ ∈ 𝐴 [𝑏 / 𝑦]πœ‘)
 
Theoremisarep2 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    &   βˆ€π‘₯ ∈ 𝐴 βˆ€π‘¦βˆ€π‘§((πœ‘ ∧ [𝑧 / 𝑦]πœ‘) β†’ 𝑦 = 𝑧)    β‡’   βˆƒπ‘€ 𝑀 = ({⟨π‘₯, π‘¦βŸ© ∣ πœ‘} β€œ 𝐴)
 
Theoremfneq1 5304 Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.)
(𝐹 = 𝐺 β†’ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴))
 
Theoremfneq2 5305 Equality theorem for function predicate with domain. (Contributed by NM, 1-Aug-1994.)
(𝐴 = 𝐡 β†’ (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐡))
 
Theoremfneq1d 5306 Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.)
(πœ‘ β†’ 𝐹 = 𝐺)    β‡’   (πœ‘ β†’ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴))
 
Theoremfneq2d 5307 Equality deduction for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.)
(πœ‘ β†’ 𝐴 = 𝐡)    β‡’   (πœ‘ β†’ (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐡))
 
Theoremfneq12d 5308 Equality deduction for function predicate with domain. (Contributed by NM, 26-Jun-2011.)
(πœ‘ β†’ 𝐹 = 𝐺)    &   (πœ‘ β†’ 𝐴 = 𝐡)    β‡’   (πœ‘ β†’ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐡))
 
Theoremfneq12 5309 Equality theorem for function predicate with domain. (Contributed by Thierry Arnoux, 31-Jan-2017.)
((𝐹 = 𝐺 ∧ 𝐴 = 𝐡) β†’ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐡))
 
Theoremfneq1i 5310 Equality inference for function predicate with domain. (Contributed by Paul Chapman, 22-Jun-2011.)
𝐹 = 𝐺    β‡’   (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)
 
Theoremfneq2i 5311 Equality inference for function predicate with domain. (Contributed by NM, 4-Sep-2011.)
𝐴 = 𝐡    β‡’   (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐡)
 
Theoremnffn 5312 Bound-variable hypothesis builder for a function with domain. (Contributed by NM, 30-Jan-2004.)
β„²π‘₯𝐹    &   β„²π‘₯𝐴    β‡’   β„²π‘₯ 𝐹 Fn 𝐴
 
Theoremfnfun 5313 A function with domain is a function. (Contributed by NM, 1-Aug-1994.)
(𝐹 Fn 𝐴 β†’ Fun 𝐹)
 
Theoremfnrel 5314 A function with domain is a relation. (Contributed by NM, 1-Aug-1994.)
(𝐹 Fn 𝐴 β†’ Rel 𝐹)
 
Theoremfndm 5315 The domain of a function. (Contributed by NM, 2-Aug-1994.)
(𝐹 Fn 𝐴 β†’ dom 𝐹 = 𝐴)
 
Theoremfunfni 5316 Inference to convert a function and domain antecedent. (Contributed by NM, 22-Apr-2004.)
((Fun 𝐹 ∧ 𝐡 ∈ dom 𝐹) β†’ πœ‘)    β‡’   ((𝐹 Fn 𝐴 ∧ 𝐡 ∈ 𝐴) β†’ πœ‘)
 
Theoremfndmu 5317 A function has a unique domain. (Contributed by NM, 11-Aug-1994.)
((𝐹 Fn 𝐴 ∧ 𝐹 Fn 𝐡) β†’ 𝐴 = 𝐡)
 
Theoremfnbr 5318 The first argument of binary relation on a function belongs to the function's domain. (Contributed by NM, 7-May-2004.)
((𝐹 Fn 𝐴 ∧ 𝐡𝐹𝐢) β†’ 𝐡 ∈ 𝐴)
 
Theoremfnop 5319 The first argument of an ordered pair in a function belongs to the function's domain. (Contributed by NM, 8-Aug-1994.)
((𝐹 Fn 𝐴 ∧ ⟨𝐡, 𝐢⟩ ∈ 𝐹) β†’ 𝐡 ∈ 𝐴)
 
Theoremfneu 5320* There is exactly one value of a function. (Contributed by NM, 22-Apr-2004.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
((𝐹 Fn 𝐴 ∧ 𝐡 ∈ 𝐴) β†’ βˆƒ!𝑦 𝐡𝐹𝑦)
 
Theoremfneu2 5321* There is exactly one value of a function. (Contributed by NM, 7-Nov-1995.)
((𝐹 Fn 𝐴 ∧ 𝐡 ∈ 𝐴) β†’ βˆƒ!π‘¦βŸ¨π΅, π‘¦βŸ© ∈ 𝐹)
 
Theoremfnun 5322 The union of two functions with disjoint domains. (Contributed by NM, 22-Sep-2004.)
(((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐡) ∧ (𝐴 ∩ 𝐡) = βˆ…) β†’ (𝐹 βˆͺ 𝐺) Fn (𝐴 βˆͺ 𝐡))
 
Theoremfnunsn 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 𝐸)
 
Theoremfnco 5324 Composition of two functions. (Contributed by NM, 22-May-2006.)
((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐡 ∧ ran 𝐺 βŠ† 𝐴) β†’ (𝐹 ∘ 𝐺) Fn 𝐡)
 
Theoremfnresdm 5325 A function does not change when restricted to its domain. (Contributed by NM, 5-Sep-2004.)
(𝐹 Fn 𝐴 β†’ (𝐹 β†Ύ 𝐴) = 𝐹)
 
Theoremfnresdisj 5326 A function restricted to a class disjoint with its domain is empty. (Contributed by NM, 23-Sep-2004.)
(𝐹 Fn 𝐴 β†’ ((𝐴 ∩ 𝐡) = βˆ… ↔ (𝐹 β†Ύ 𝐡) = βˆ…))
 
Theorem2elresin 5327 Membership in two functions restricted by each other's domain. (Contributed by NM, 8-Aug-1994.)
((𝐹 Fn 𝐴 ∧ 𝐺 Fn 𝐡) β†’ ((⟨π‘₯, π‘¦βŸ© ∈ 𝐹 ∧ ⟨π‘₯, π‘§βŸ© ∈ 𝐺) ↔ (⟨π‘₯, π‘¦βŸ© ∈ (𝐹 β†Ύ (𝐴 ∩ 𝐡)) ∧ ⟨π‘₯, π‘§βŸ© ∈ (𝐺 β†Ύ (𝐴 ∩ 𝐡)))))
 
Theoremfnssresb 5328 Restriction of a function with a subclass of its domain. (Contributed by NM, 10-Oct-2007.)
(𝐹 Fn 𝐴 β†’ ((𝐹 β†Ύ 𝐡) Fn 𝐡 ↔ 𝐡 βŠ† 𝐴))
 
Theoremfnssres 5329 Restriction of a function with a subclass of its domain. (Contributed by NM, 2-Aug-1994.)
((𝐹 Fn 𝐴 ∧ 𝐡 βŠ† 𝐴) β†’ (𝐹 β†Ύ 𝐡) Fn 𝐡)
 
Theoremfnresin1 5330 Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.)
(𝐹 Fn 𝐴 β†’ (𝐹 β†Ύ (𝐴 ∩ 𝐡)) Fn (𝐴 ∩ 𝐡))
 
Theoremfnresin2 5331 Restriction of a function's domain with an intersection. (Contributed by NM, 9-Aug-1994.)
(𝐹 Fn 𝐴 β†’ (𝐹 β†Ύ (𝐡 ∩ 𝐴)) Fn (𝐡 ∩ 𝐴))
 
Theoremfnres 5332* An equivalence for functionality of a restriction. Compare dffun8 5244. (Contributed by Mario Carneiro, 20-May-2015.)
((𝐹 β†Ύ 𝐴) Fn 𝐴 ↔ βˆ€π‘₯ ∈ 𝐴 βˆƒ!𝑦 π‘₯𝐹𝑦)
 
Theoremfnresi 5333 Functionality and domain of restricted identity. (Contributed by NM, 27-Aug-2004.)
( I β†Ύ 𝐴) Fn 𝐴
 
Theoremfnima 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 𝐹)
 
Theoremfn0 5335 A function with empty domain is empty. (Contributed by NM, 15-Apr-1998.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
(𝐹 Fn βˆ… ↔ 𝐹 = βˆ…)
 
Theoremfnimadisj 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 𝐴 ∧ (𝐴 ∩ 𝐢) = βˆ…) β†’ (𝐹 β€œ 𝐢) = βˆ…)
 
Theoremfnimaeq0 5337 Images under a function never map nonempty sets to empty sets. (Contributed by Stefan O'Rear, 21-Jan-2015.)
((𝐹 Fn 𝐴 ∧ 𝐡 βŠ† 𝐴) β†’ ((𝐹 β€œ 𝐡) = βˆ… ↔ 𝐡 = βˆ…))
 
Theoremdfmpt3 5338 Alternate definition for the maps-to notation df-mpt 4066. (Contributed by Mario Carneiro, 30-Dec-2016.)
(π‘₯ ∈ 𝐴 ↦ 𝐡) = βˆͺ π‘₯ ∈ 𝐴 ({π‘₯} Γ— {𝐡})
 
Theoremfnopabg 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 𝐴)
 
Theoremfnopab 5340* Functionality and domain of an ordered-pair class abstraction. (Contributed by NM, 5-Mar-1996.)
(π‘₯ ∈ 𝐴 β†’ βˆƒ!π‘¦πœ‘)    &   πΉ = {⟨π‘₯, π‘¦βŸ© ∣ (π‘₯ ∈ 𝐴 ∧ πœ‘)}    β‡’   πΉ Fn 𝐴
 
Theoremmptfng 5341* The maps-to notation defines a function with domain. (Contributed by Scott Fenton, 21-Mar-2011.)
𝐹 = (π‘₯ ∈ 𝐴 ↦ 𝐡)    β‡’   (βˆ€π‘₯ ∈ 𝐴 𝐡 ∈ V ↔ 𝐹 Fn 𝐴)
 
Theoremfnmpt 5342* The maps-to notation defines a function with domain. (Contributed by NM, 9-Apr-2013.)
𝐹 = (π‘₯ ∈ 𝐴 ↦ 𝐡)    β‡’   (βˆ€π‘₯ ∈ 𝐴 𝐡 ∈ 𝑉 β†’ 𝐹 Fn 𝐴)
 
Theoremmpt0 5343 A mapping operation with empty domain. (Contributed by Mario Carneiro, 28-Dec-2014.)
(π‘₯ ∈ βˆ… ↦ 𝐴) = βˆ…
 
Theoremfnmpti 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 𝐴
 
Theoremdmmpti 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 𝐹 = 𝐴
 
Theoremdmmptd 5346* The domain of the mapping operation, deduction form. (Contributed by Glauco Siliprandi, 11-Dec-2019.)
𝐴 = (π‘₯ ∈ 𝐡 ↦ 𝐢)    &   ((πœ‘ ∧ π‘₯ ∈ 𝐡) β†’ 𝐢 ∈ 𝑉)    β‡’   (πœ‘ β†’ dom 𝐴 = 𝐡)
 
Theoremmptun 5347 Union of mappings which are mutually compatible. (Contributed by Mario Carneiro, 31-Aug-2015.)
(π‘₯ ∈ (𝐴 βˆͺ 𝐡) ↦ 𝐢) = ((π‘₯ ∈ 𝐴 ↦ 𝐢) βˆͺ (π‘₯ ∈ 𝐡 ↦ 𝐢))
 
Theoremfeq1 5348 Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
(𝐹 = 𝐺 β†’ (𝐹:𝐴⟢𝐡 ↔ 𝐺:𝐴⟢𝐡))
 
Theoremfeq2 5349 Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
(𝐴 = 𝐡 β†’ (𝐹:𝐴⟢𝐢 ↔ 𝐹:𝐡⟢𝐢))
 
Theoremfeq3 5350 Equality theorem for functions. (Contributed by NM, 1-Aug-1994.)
(𝐴 = 𝐡 β†’ (𝐹:𝐢⟢𝐴 ↔ 𝐹:𝐢⟢𝐡))
 
Theoremfeq23 5351 Equality theorem for functions. (Contributed by FL, 14-Jul-2007.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
((𝐴 = 𝐢 ∧ 𝐡 = 𝐷) β†’ (𝐹:𝐴⟢𝐡 ↔ 𝐹:𝐢⟢𝐷))
 
Theoremfeq1d 5352 Equality deduction for functions. (Contributed by NM, 19-Feb-2008.)
(πœ‘ β†’ 𝐹 = 𝐺)    β‡’   (πœ‘ β†’ (𝐹:𝐴⟢𝐡 ↔ 𝐺:𝐴⟢𝐡))
 
Theoremfeq2d 5353 Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.)
(πœ‘ β†’ 𝐴 = 𝐡)    β‡’   (πœ‘ β†’ (𝐹:𝐴⟢𝐢 ↔ 𝐹:𝐡⟢𝐢))
 
Theoremfeq3d 5354 Equality deduction for functions. (Contributed by AV, 1-Jan-2020.)
(πœ‘ β†’ 𝐴 = 𝐡)    β‡’   (πœ‘ β†’ (𝐹:π‘‹βŸΆπ΄ ↔ 𝐹:π‘‹βŸΆπ΅))
 
Theoremfeq12d 5355 Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.)
(πœ‘ β†’ 𝐹 = 𝐺)    &   (πœ‘ β†’ 𝐴 = 𝐡)    β‡’   (πœ‘ β†’ (𝐹:𝐴⟢𝐢 ↔ 𝐺:𝐡⟢𝐢))
 
Theoremfeq123d 5356 Equality deduction for functions. (Contributed by Paul Chapman, 22-Jun-2011.)
(πœ‘ β†’ 𝐹 = 𝐺)    &   (πœ‘ β†’ 𝐴 = 𝐡)    &   (πœ‘ β†’ 𝐢 = 𝐷)    β‡’   (πœ‘ β†’ (𝐹:𝐴⟢𝐢 ↔ 𝐺:𝐡⟢𝐷))
 
Theoremfeq123 5357 Equality theorem for functions. (Contributed by FL, 16-Nov-2008.)
((𝐹 = 𝐺 ∧ 𝐴 = 𝐢 ∧ 𝐡 = 𝐷) β†’ (𝐹:𝐴⟢𝐡 ↔ 𝐺:𝐢⟢𝐷))
 
Theoremfeq1i 5358 Equality inference for functions. (Contributed by Paul Chapman, 22-Jun-2011.)
𝐹 = 𝐺    β‡’   (𝐹:𝐴⟢𝐡 ↔ 𝐺:𝐴⟢𝐡)
 
Theoremfeq2i 5359 Equality inference for functions. (Contributed by NM, 5-Sep-2011.)
𝐴 = 𝐡    β‡’   (𝐹:𝐴⟢𝐢 ↔ 𝐹:𝐡⟢𝐢)
 
Theoremfeq23i 5360 Equality inference for functions. (Contributed by Paul Chapman, 22-Jun-2011.)
𝐴 = 𝐢    &   π΅ = 𝐷    β‡’   (𝐹:𝐴⟢𝐡 ↔ 𝐹:𝐢⟢𝐷)
 
Theoremfeq23d 5361 Equality deduction for functions. (Contributed by NM, 8-Jun-2013.)
(πœ‘ β†’ 𝐴 = 𝐢)    &   (πœ‘ β†’ 𝐡 = 𝐷)    β‡’   (πœ‘ β†’ (𝐹:𝐴⟢𝐡 ↔ 𝐹:𝐢⟢𝐷))
 
Theoremnff 5362 Bound-variable hypothesis builder for a mapping. (Contributed by NM, 29-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.)
β„²π‘₯𝐹    &   β„²π‘₯𝐴    &   β„²π‘₯𝐡    β‡’   β„²π‘₯ 𝐹:𝐴⟢𝐡
 
Theoremsbcfng 5363* Distribute proper substitution through the function predicate with a domain. (Contributed by Alexander van der Vekens, 15-Jul-2018.)
(𝑋 ∈ 𝑉 β†’ ([𝑋 / π‘₯]𝐹 Fn 𝐴 ↔ ⦋𝑋 / π‘₯⦌𝐹 Fn ⦋𝑋 / π‘₯⦌𝐴))
 
Theoremsbcfg 5364* Distribute proper substitution through the function predicate with domain and codomain. (Contributed by Alexander van der Vekens, 15-Jul-2018.)
(𝑋 ∈ 𝑉 β†’ ([𝑋 / π‘₯]𝐹:𝐴⟢𝐡 ↔ ⦋𝑋 / π‘₯⦌𝐹:⦋𝑋 / π‘₯β¦Œπ΄βŸΆβ¦‹π‘‹ / π‘₯⦌𝐡))
 
Theoremffn 5365 A mapping is a function. (Contributed by NM, 2-Aug-1994.)
(𝐹:𝐴⟢𝐡 β†’ 𝐹 Fn 𝐴)
 
Theoremffnd 5366 A mapping is a function with domain, deduction form. (Contributed by Glauco Siliprandi, 17-Aug-2020.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    β‡’   (πœ‘ β†’ 𝐹 Fn 𝐴)
 
Theoremdffn2 5367 Any function is a mapping into V. (Contributed by NM, 31-Oct-1995.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
(𝐹 Fn 𝐴 ↔ 𝐹:𝐴⟢V)
 
Theoremffun 5368 A mapping is a function. (Contributed by NM, 3-Aug-1994.)
(𝐹:𝐴⟢𝐡 β†’ Fun 𝐹)
 
Theoremffund 5369 A mapping is a function, deduction version. (Contributed by Glauco Siliprandi, 3-Mar-2021.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    β‡’   (πœ‘ β†’ Fun 𝐹)
 
Theoremfrel 5370 A mapping is a relation. (Contributed by NM, 3-Aug-1994.)
(𝐹:𝐴⟢𝐡 β†’ Rel 𝐹)
 
Theoremfdm 5371 The domain of a mapping. (Contributed by NM, 2-Aug-1994.)
(𝐹:𝐴⟢𝐡 β†’ dom 𝐹 = 𝐴)
 
Theoremfdmd 5372 Deduction form of fdm 5371. The domain of a mapping. (Contributed by Glauco Siliprandi, 26-Jun-2021.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    β‡’   (πœ‘ β†’ dom 𝐹 = 𝐴)
 
Theoremfdmi 5373 The domain of a mapping. (Contributed by NM, 28-Jul-2008.)
𝐹:𝐴⟢𝐡    β‡’   dom 𝐹 = 𝐴
 
Theoremfrn 5374 The range of a mapping. (Contributed by NM, 3-Aug-1994.)
(𝐹:𝐴⟢𝐡 β†’ ran 𝐹 βŠ† 𝐡)
 
Theoremfrnd 5375 Deduction form of frn 5374. The range of a mapping. (Contributed by Glauco Siliprandi, 26-Jun-2021.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    β‡’   (πœ‘ β†’ ran 𝐹 βŠ† 𝐡)
 
Theoremdffn3 5376 A function maps to its range. (Contributed by NM, 1-Sep-1999.)
(𝐹 Fn 𝐴 ↔ 𝐹:𝐴⟢ran 𝐹)
 
Theoremfss 5377 Expanding the codomain of a mapping. (Contributed by NM, 10-May-1998.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
((𝐹:𝐴⟢𝐡 ∧ 𝐡 βŠ† 𝐢) β†’ 𝐹:𝐴⟢𝐢)
 
Theoremfssd 5378 Expanding the codomain of a mapping, deduction form. (Contributed by Glauco Siliprandi, 11-Dec-2019.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    &   (πœ‘ β†’ 𝐡 βŠ† 𝐢)    β‡’   (πœ‘ β†’ 𝐹:𝐴⟢𝐢)
 
Theoremfssdmd 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 𝐹)    β‡’   (πœ‘ β†’ 𝐷 βŠ† 𝐴)
 
Theoremfssdm 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 𝐹    &   (πœ‘ β†’ 𝐹:𝐴⟢𝐡)    β‡’   (πœ‘ β†’ 𝐷 βŠ† 𝐴)
 
Theoremfco 5381 Composition of two mappings. (Contributed by NM, 29-Aug-1999.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
((𝐹:𝐡⟢𝐢 ∧ 𝐺:𝐴⟢𝐡) β†’ (𝐹 ∘ 𝐺):𝐴⟢𝐢)
 
Theoremfco2 5382 Functionality of a composition with weakened out of domain condition on the first argument. (Contributed by Stefan O'Rear, 11-Mar-2015.)
(((𝐹 β†Ύ 𝐡):𝐡⟢𝐢 ∧ 𝐺:𝐴⟢𝐡) β†’ (𝐹 ∘ 𝐺):𝐴⟢𝐢)
 
Theoremfssxp 5383 A mapping is a class of ordered pairs. (Contributed by NM, 3-Aug-1994.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
(𝐹:𝐴⟢𝐡 β†’ 𝐹 βŠ† (𝐴 Γ— 𝐡))
 
Theoremfex2 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)
 
Theoremfunssxp 5385 Two ways of specifying a partial function from 𝐴 to 𝐡. (Contributed by NM, 13-Nov-2007.)
((Fun 𝐹 ∧ 𝐹 βŠ† (𝐴 Γ— 𝐡)) ↔ (𝐹:dom 𝐹⟢𝐡 ∧ dom 𝐹 βŠ† 𝐴))
 
Theoremffdm 5386 A mapping is a partial function. (Contributed by NM, 25-Nov-2007.)
(𝐹:𝐴⟢𝐡 β†’ (𝐹:dom 𝐹⟢𝐡 ∧ dom 𝐹 βŠ† 𝐴))
 
Theoremopelf 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.)
((𝐹:𝐴⟢𝐡 ∧ ⟨𝐢, 𝐷⟩ ∈ 𝐹) β†’ (𝐢 ∈ 𝐴 ∧ 𝐷 ∈ 𝐡))
 
Theoremfun 5388 The union of two functions with disjoint domains. (Contributed by NM, 22-Sep-2004.)
(((𝐹:𝐴⟢𝐢 ∧ 𝐺:𝐡⟢𝐷) ∧ (𝐴 ∩ 𝐡) = βˆ…) β†’ (𝐹 βˆͺ 𝐺):(𝐴 βˆͺ 𝐡)⟢(𝐢 βˆͺ 𝐷))
 
Theoremfun2 5389 The union of two functions with disjoint domains. (Contributed by Mario Carneiro, 12-Mar-2015.)
(((𝐹:𝐴⟢𝐢 ∧ 𝐺:𝐡⟢𝐢) ∧ (𝐴 ∩ 𝐡) = βˆ…) β†’ (𝐹 βˆͺ 𝐺):(𝐴 βˆͺ 𝐡)⟢𝐢)
 
Theoremfnfco 5390 Composition of two functions. (Contributed by NM, 22-May-2006.)
((𝐹 Fn 𝐴 ∧ 𝐺:𝐡⟢𝐴) β†’ (𝐹 ∘ 𝐺) Fn 𝐡)
 
Theoremfssres 5391 Restriction of a function with a subclass of its domain. (Contributed by NM, 23-Sep-2004.)
((𝐹:𝐴⟢𝐡 ∧ 𝐢 βŠ† 𝐴) β†’ (𝐹 β†Ύ 𝐢):𝐢⟢𝐡)
 
Theoremfssresd 5392 Restriction of a function with a subclass of its domain, deduction form. (Contributed by Glauco Siliprandi, 11-Dec-2019.)
(πœ‘ β†’ 𝐹:𝐴⟢𝐡)    &   (πœ‘ β†’ 𝐢 βŠ† 𝐴)    β‡’   (πœ‘ β†’ (𝐹 β†Ύ 𝐢):𝐢⟢𝐡)
 
Theoremfssres2 5393 Restriction of a restricted function with a subclass of its domain. (Contributed by NM, 21-Jul-2005.)
(((𝐹 β†Ύ 𝐴):𝐴⟢𝐡 ∧ 𝐢 βŠ† 𝐴) β†’ (𝐹 β†Ύ 𝐢):𝐢⟢𝐡)
 
Theoremfresin 5394 An identity for the mapping relationship under restriction. (Contributed by Scott Fenton, 4-Sep-2011.) (Proof shortened by Mario Carneiro, 26-May-2016.)
(𝐹:𝐴⟢𝐡 β†’ (𝐹 β†Ύ 𝑋):(𝐴 ∩ 𝑋)⟢𝐡)
 
Theoremresasplitss 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 𝐡 ∧ (𝐹 β†Ύ (𝐴 ∩ 𝐡)) = (𝐺 β†Ύ (𝐴 ∩ 𝐡))) β†’ ((𝐹 β†Ύ (𝐴 ∩ 𝐡)) βˆͺ ((𝐹 β†Ύ (𝐴 βˆ– 𝐡)) βˆͺ (𝐺 β†Ύ (𝐡 βˆ– 𝐴)))) βŠ† (𝐹 βˆͺ 𝐺))
 
Theoremfcoi1 5396 Composition of a mapping and restricted identity. (Contributed by NM, 13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
(𝐹:𝐴⟢𝐡 β†’ (𝐹 ∘ ( I β†Ύ 𝐴)) = 𝐹)
 
Theoremfcoi2 5397 Composition of restricted identity and a mapping. (Contributed by NM, 13-Dec-2003.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
(𝐹:𝐴⟢𝐡 β†’ (( I β†Ύ 𝐡) ∘ 𝐹) = 𝐹)
 
Theoremfeu 5398* There is exactly one value of a function in its codomain. (Contributed by NM, 10-Dec-2003.)
((𝐹:𝐴⟢𝐡 ∧ 𝐢 ∈ 𝐴) β†’ βˆƒ!𝑦 ∈ 𝐡 ⟨𝐢, π‘¦βŸ© ∈ 𝐹)
 
Theoremfcnvres 5399 The converse of a restriction of a function. (Contributed by NM, 26-Mar-1998.)
(𝐹:𝐴⟢𝐡 β†’ β—‘(𝐹 β†Ύ 𝐴) = (◑𝐹 β†Ύ 𝐡))
 
Theoremfimacnvdisj 5400 The preimage of a class disjoint with a mapping's codomain is empty. (Contributed by FL, 24-Jan-2007.)
((𝐹:𝐴⟢𝐡 ∧ (𝐡 ∩ 𝐢) = βˆ…) β†’ (◑𝐹 β€œ 𝐢) = βˆ…)
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10200 103 10201-10300 104 10301-10400 105 10401-10500 106 10501-10600 107 10601-10700 108 10701-10800 109 10801-10900 110 10901-11000 111 11001-11100 112 11101-11200 113 11201-11300 114 11301-11400 115 11401-11500 116 11501-11600 117 11601-11700 118 11701-11800 119 11801-11900 120 11901-12000 121 12001-12100 122 12101-12200 123 12201-12300 124 12301-12400 125 12401-12500 126 12501-12600 127 12601-12700 128 12701-12800 129 12801-12900 130 12901-13000 131 13001-13100 132 13101-13200 133 13201-13300 134 13301-13400 135 13401-13500 136 13501-13600 137 13601-13700 138 13701-13800 139 13801-13900 140 13901-14000 141 14001-14100 142 14101-14200 143 14201-14300 144 14301-14400 145 14401-14500 146 14501-14600 147 14601-14700 148 14701-14800 149 14801-14801
  Copyright terms: Public domain < Previous  Next >