![]() |
Intuitionistic Logic Explorer Theorem List (p. 53 of 153) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | cbviota 5201 | Change bound variables in a description binder. (Contributed by Andrew Salmon, 1-Aug-2011.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 ⇒ ⊢ (℩𝑥𝜑) = (℩𝑦𝜓) | ||
Theorem | cbviotav 5202* | Change bound variables in a description binder. (Contributed by Andrew Salmon, 1-Aug-2011.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (℩𝑥𝜑) = (℩𝑦𝜓) | ||
Theorem | sb8iota 5203 | Variable substitution in description binder. Compare sb8eu 2051. (Contributed by NM, 18-Mar-2013.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (℩𝑥𝜑) = (℩𝑦[𝑦 / 𝑥]𝜑) | ||
Theorem | iotaeq 5204 | Equality theorem for descriptions. (Contributed by Andrew Salmon, 30-Jun-2011.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (℩𝑥𝜑) = (℩𝑦𝜑)) | ||
Theorem | iotabi 5205 | Equivalence theorem for descriptions. (Contributed by Andrew Salmon, 30-Jun-2011.) |
⊢ (∀𝑥(𝜑 ↔ 𝜓) → (℩𝑥𝜑) = (℩𝑥𝜓)) | ||
Theorem | uniabio 5206* | Part of Theorem 8.17 in [Quine] p. 56. This theorem serves as a lemma for the fundamental property of iota. (Contributed by Andrew Salmon, 11-Jul-2011.) |
⊢ (∀𝑥(𝜑 ↔ 𝑥 = 𝑦) → ∪ {𝑥 ∣ 𝜑} = 𝑦) | ||
Theorem | iotaval 5207* | Theorem 8.19 in [Quine] p. 57. This theorem is the fundamental property of iota. (Contributed by Andrew Salmon, 11-Jul-2011.) |
⊢ (∀𝑥(𝜑 ↔ 𝑥 = 𝑦) → (℩𝑥𝜑) = 𝑦) | ||
Theorem | iotauni 5208 | Equivalence between two different forms of ℩. (Contributed by Andrew Salmon, 12-Jul-2011.) |
⊢ (∃!𝑥𝜑 → (℩𝑥𝜑) = ∪ {𝑥 ∣ 𝜑}) | ||
Theorem | iotaint 5209 | Equivalence between two different forms of ℩. (Contributed by Mario Carneiro, 24-Dec-2016.) |
⊢ (∃!𝑥𝜑 → (℩𝑥𝜑) = ∩ {𝑥 ∣ 𝜑}) | ||
Theorem | iota1 5210 | Property of iota. (Contributed by NM, 23-Aug-2011.) (Revised by Mario Carneiro, 23-Dec-2016.) |
⊢ (∃!𝑥𝜑 → (𝜑 ↔ (℩𝑥𝜑) = 𝑥)) | ||
Theorem | iotanul 5211 | Theorem 8.22 in [Quine] p. 57. This theorem is the result if there isn't exactly one 𝑥 that satisfies 𝜑. (Contributed by Andrew Salmon, 11-Jul-2011.) |
⊢ (¬ ∃!𝑥𝜑 → (℩𝑥𝜑) = ∅) | ||
Theorem | euiotaex 5212 | Theorem 8.23 in [Quine] p. 58, with existential uniqueness condition added. This theorem proves the existence of the ℩ class under our definition. (Contributed by Jim Kingdon, 21-Dec-2018.) |
⊢ (∃!𝑥𝜑 → (℩𝑥𝜑) ∈ V) | ||
Theorem | iotass 5213* | Value of iota based on a proposition which holds only for values which are subsets of a given class. (Contributed by Mario Carneiro and Jim Kingdon, 21-Dec-2018.) |
⊢ (∀𝑥(𝜑 → 𝑥 ⊆ 𝐴) → (℩𝑥𝜑) ⊆ 𝐴) | ||
Theorem | iotaexab 5214 | Existence of the ℩ class when all the possible values are contained in a set. (Contributed by Jim Kingdon, 27-May-2025.) |
⊢ ({𝑥 ∣ 𝜑} ∈ 𝑉 → (℩𝑥𝜑) ∈ V) | ||
Theorem | iota4 5215 | Theorem *14.22 in [WhiteheadRussell] p. 190. (Contributed by Andrew Salmon, 12-Jul-2011.) |
⊢ (∃!𝑥𝜑 → [(℩𝑥𝜑) / 𝑥]𝜑) | ||
Theorem | iota4an 5216 | Theorem *14.23 in [WhiteheadRussell] p. 191. (Contributed by Andrew Salmon, 12-Jul-2011.) |
⊢ (∃!𝑥(𝜑 ∧ 𝜓) → [(℩𝑥(𝜑 ∧ 𝜓)) / 𝑥]𝜑) | ||
Theorem | iota5 5217* | A method for computing iota. (Contributed by NM, 17-Sep-2013.) |
⊢ ((𝜑 ∧ 𝐴 ∈ 𝑉) → (𝜓 ↔ 𝑥 = 𝐴)) ⇒ ⊢ ((𝜑 ∧ 𝐴 ∈ 𝑉) → (℩𝑥𝜓) = 𝐴) | ||
Theorem | iotabidv 5218* | Formula-building deduction for iota. (Contributed by NM, 20-Aug-2011.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (℩𝑥𝜓) = (℩𝑥𝜒)) | ||
Theorem | iotabii 5219 | Formula-building deduction for iota. (Contributed by Mario Carneiro, 2-Oct-2015.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (℩𝑥𝜑) = (℩𝑥𝜓) | ||
Theorem | iotacl 5220 |
Membership law for descriptions.
This can useful for expanding an unbounded iota-based definition (see df-iota 5196). (Contributed by Andrew Salmon, 1-Aug-2011.) |
⊢ (∃!𝑥𝜑 → (℩𝑥𝜑) ∈ {𝑥 ∣ 𝜑}) | ||
Theorem | iota2df 5221 | A condition that allows us to represent "the unique element such that 𝜑 " with a class expression 𝐴. (Contributed by NM, 30-Dec-2014.) |
⊢ (𝜑 → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → ∃!𝑥𝜓) & ⊢ ((𝜑 ∧ 𝑥 = 𝐵) → (𝜓 ↔ 𝜒)) & ⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → Ⅎ𝑥𝐵) ⇒ ⊢ (𝜑 → (𝜒 ↔ (℩𝑥𝜓) = 𝐵)) | ||
Theorem | iota2d 5222* | A condition that allows us to represent "the unique element such that 𝜑 " with a class expression 𝐴. (Contributed by NM, 30-Dec-2014.) |
⊢ (𝜑 → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → ∃!𝑥𝜓) & ⊢ ((𝜑 ∧ 𝑥 = 𝐵) → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜒 ↔ (℩𝑥𝜓) = 𝐵)) | ||
Theorem | eliota 5223* | An element of an iota expression. (Contributed by Jim Kingdon, 22-Nov-2024.) |
⊢ (𝐴 ∈ (℩𝑥𝜑) ↔ ∃𝑦(𝐴 ∈ 𝑦 ∧ ∀𝑥(𝜑 ↔ 𝑥 = 𝑦))) | ||
Theorem | eliotaeu 5224 | An inhabited iota expression has a unique value. (Contributed by Jim Kingdon, 22-Nov-2024.) |
⊢ (𝐴 ∈ (℩𝑥𝜑) → ∃!𝑥𝜑) | ||
Theorem | iota2 5225* | The unique element such that 𝜑. (Contributed by Jeff Madsen, 1-Jun-2011.) (Revised by Mario Carneiro, 23-Dec-2016.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐴 ∈ 𝐵 ∧ ∃!𝑥𝜑) → (𝜓 ↔ (℩𝑥𝜑) = 𝐴)) | ||
Theorem | sniota 5226 | A class abstraction with a unique member can be expressed as a singleton. (Contributed by Mario Carneiro, 23-Dec-2016.) |
⊢ (∃!𝑥𝜑 → {𝑥 ∣ 𝜑} = {(℩𝑥𝜑)}) | ||
Theorem | iotam 5227* | Representation of "the unique element such that 𝜑 " with a class expression 𝐴 which is inhabited (that means that "the unique element such that 𝜑 " exists). (Contributed by AV, 30-Jan-2024.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐴 ∈ 𝑉 ∧ ∃𝑤 𝑤 ∈ 𝐴 ∧ 𝐴 = (℩𝑥𝜑)) → 𝜓) | ||
Theorem | csbiotag 5228* | Class substitution within a description binder. (Contributed by Scott Fenton, 6-Oct-2017.) |
⊢ (𝐴 ∈ 𝑉 → ⦋𝐴 / 𝑥⦌(℩𝑦𝜑) = (℩𝑦[𝐴 / 𝑥]𝜑)) | ||
Syntax | wfun 5229 | Extend the definition of a wff to include the function predicate. (Read: 𝐴 is a function.) |
wff Fun 𝐴 | ||
Syntax | wfn 5230 | Extend the definition of a wff to include the function predicate with a domain. (Read: 𝐴 is a function on 𝐵.) |
wff 𝐴 Fn 𝐵 | ||
Syntax | wf 5231 | Extend the definition of a wff to include the function predicate with domain and codomain. (Read: 𝐹 maps 𝐴 into 𝐵.) |
wff 𝐹:𝐴⟶𝐵 | ||
Syntax | wf1 5232 | Extend the definition of a wff to include one-to-one functions. (Read: 𝐹 maps 𝐴 one-to-one into 𝐵.) The notation ("1-1" above the arrow) is from Definition 6.15(5) of [TakeutiZaring] p. 27. |
wff 𝐹:𝐴–1-1→𝐵 | ||
Syntax | wfo 5233 | Extend the definition of a wff to include onto functions. (Read: 𝐹 maps 𝐴 onto 𝐵.) The notation ("onto" below the arrow) is from Definition 6.15(4) of [TakeutiZaring] p. 27. |
wff 𝐹:𝐴–onto→𝐵 | ||
Syntax | wf1o 5234 | Extend the definition of a wff to include one-to-one onto functions. (Read: 𝐹 maps 𝐴 one-to-one onto 𝐵.) The notation ("1-1" above the arrow and "onto" below the arrow) is from Definition 6.15(6) of [TakeutiZaring] p. 27. |
wff 𝐹:𝐴–1-1-onto→𝐵 | ||
Syntax | cfv 5235 | Extend the definition of a class to include the value of a function. (Read: The value of 𝐹 at 𝐴, or "𝐹 of 𝐴.") |
class (𝐹‘𝐴) | ||
Syntax | wiso 5236 | Extend the definition of a wff to include the isomorphism property. (Read: 𝐻 is an 𝑅, 𝑆 isomorphism of 𝐴 onto 𝐵.) |
wff 𝐻 Isom 𝑅, 𝑆 (𝐴, 𝐵) | ||
Definition | df-fun 5237 | Define predicate that determines if some class 𝐴 is a function. Definition 10.1 of [Quine] p. 65. For example, the expression Fun I is true (funi 5267). This is not the same as defining a specific function's mapping, which is typically done using the format of cmpt 4079 with the maps-to notation (see df-mpt 4081). Contrast this predicate with the predicates to determine if some class is a function with a given domain (df-fn 5238), a function with a given domain and codomain (df-f 5239), a one-to-one function (df-f1 5240), an onto function (df-fo 5241), or a one-to-one onto function (df-f1o 5242). For alternate definitions, see dffun2 5245, dffun4 5246, dffun6 5249, dffun7 5262, dffun8 5263, and dffun9 5264. (Contributed by NM, 1-Aug-1994.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ (𝐴 ∘ ◡𝐴) ⊆ I )) | ||
Definition | df-fn 5238 | Define a function with domain. Definition 6.15(1) of [TakeutiZaring] p. 27. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐴 Fn 𝐵 ↔ (Fun 𝐴 ∧ dom 𝐴 = 𝐵)) | ||
Definition | df-f 5239 | Define a function (mapping) with domain and codomain. Definition 6.15(3) of [TakeutiZaring] p. 27. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹:𝐴⟶𝐵 ↔ (𝐹 Fn 𝐴 ∧ ran 𝐹 ⊆ 𝐵)) | ||
Definition | df-f1 5240 | Define a one-to-one function. Compare Definition 6.15(5) of [TakeutiZaring] p. 27. We use their notation ("1-1" above the arrow). (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹:𝐴–1-1→𝐵 ↔ (𝐹:𝐴⟶𝐵 ∧ Fun ◡𝐹)) | ||
Definition | df-fo 5241 | Define an onto function. Definition 6.15(4) of [TakeutiZaring] p. 27. We use their notation ("onto" under the arrow). (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹:𝐴–onto→𝐵 ↔ (𝐹 Fn 𝐴 ∧ ran 𝐹 = 𝐵)) | ||
Definition | df-f1o 5242 | Define a one-to-one onto function. Compare Definition 6.15(6) of [TakeutiZaring] p. 27. We use their notation ("1-1" above the arrow and "onto" below the arrow). (Contributed by NM, 1-Aug-1994.) |
⊢ (𝐹:𝐴–1-1-onto→𝐵 ↔ (𝐹:𝐴–1-1→𝐵 ∧ 𝐹:𝐴–onto→𝐵)) | ||
Definition | df-fv 5243* | Define the value of a function, (𝐹‘𝐴), also known as function application. For example, ( I ‘∅) = ∅. Typically, function 𝐹 is defined using maps-to notation (see df-mpt 4081), but this is not required. For example, 𝐹 = {〈2, 6〉, 〈3, 9〉} → (𝐹‘3) = 9. We will later define two-argument functions using ordered pairs as (𝐴𝐹𝐵) = (𝐹‘〈𝐴, 𝐵〉). This particular definition is quite convenient: it can be applied to any class and evaluates to the empty set when it is not meaningful. The left apostrophe notation originated with Peano and was adopted in Definition *30.01 of [WhiteheadRussell] p. 235, Definition 10.11 of [Quine] p. 68, and Definition 6.11 of [TakeutiZaring] p. 26. It means the same thing as the more familiar 𝐹(𝐴) notation for a function's value at 𝐴, i.e., "𝐹 of 𝐴," but without context-dependent notational ambiguity. (Contributed by NM, 1-Aug-1994.) Revised to use ℩. (Revised by Scott Fenton, 6-Oct-2017.) |
⊢ (𝐹‘𝐴) = (℩𝑥𝐴𝐹𝑥) | ||
Definition | df-isom 5244* | Define the isomorphism predicate. We read this as "𝐻 is an 𝑅, 𝑆 isomorphism of 𝐴 onto 𝐵". Normally, 𝑅 and 𝑆 are ordering relations on 𝐴 and 𝐵 respectively. Definition 6.28 of [TakeutiZaring] p. 32, whose notation is the same as ours except that 𝑅 and 𝑆 are subscripts. (Contributed by NM, 4-Mar-1997.) |
⊢ (𝐻 Isom 𝑅, 𝑆 (𝐴, 𝐵) ↔ (𝐻:𝐴–1-1-onto→𝐵 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝑥𝑅𝑦 ↔ (𝐻‘𝑥)𝑆(𝐻‘𝑦)))) | ||
Theorem | dffun2 5245* | Alternate definition of a function. (Contributed by NM, 29-Dec-1996.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((𝑥𝐴𝑦 ∧ 𝑥𝐴𝑧) → 𝑦 = 𝑧))) | ||
Theorem | dffun4 5246* | Alternate definition of a function. Definition 6.4(4) of [TakeutiZaring] p. 24. (Contributed by NM, 29-Dec-1996.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((〈𝑥, 𝑦〉 ∈ 𝐴 ∧ 〈𝑥, 𝑧〉 ∈ 𝐴) → 𝑦 = 𝑧))) | ||
Theorem | dffun5r 5247* | A way of proving a relation is a function, analogous to mo2r 2090. (Contributed by Jim Kingdon, 27-May-2020.) |
⊢ ((Rel 𝐴 ∧ ∀𝑥∃𝑧∀𝑦(〈𝑥, 𝑦〉 ∈ 𝐴 → 𝑦 = 𝑧)) → Fun 𝐴) | ||
Theorem | dffun6f 5248* | Definition of function, using bound-variable hypotheses instead of distinct variable conditions. (Contributed by NM, 9-Mar-1995.) (Revised by Mario Carneiro, 15-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑦𝐴 ⇒ ⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∃*𝑦 𝑥𝐴𝑦)) | ||
Theorem | dffun6 5249* | Alternate definition of a function using "at most one" notation. (Contributed by NM, 9-Mar-1995.) |
⊢ (Fun 𝐹 ↔ (Rel 𝐹 ∧ ∀𝑥∃*𝑦 𝑥𝐹𝑦)) | ||
Theorem | funmo 5250* | A function has at most one value for each argument. (Contributed by NM, 24-May-1998.) |
⊢ (Fun 𝐹 → ∃*𝑦 𝐴𝐹𝑦) | ||
Theorem | dffun4f 5251* | Definition of function like dffun4 5246 but using bound-variable hypotheses instead of distinct variable conditions. (Contributed by Jim Kingdon, 17-Mar-2019.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑦𝐴 & ⊢ Ⅎ𝑧𝐴 ⇒ ⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((〈𝑥, 𝑦〉 ∈ 𝐴 ∧ 〈𝑥, 𝑧〉 ∈ 𝐴) → 𝑦 = 𝑧))) | ||
Theorem | funrel 5252 | A function is a relation. (Contributed by NM, 1-Aug-1994.) |
⊢ (Fun 𝐴 → Rel 𝐴) | ||
Theorem | 0nelfun 5253 | A function does not contain the empty set. (Contributed by BJ, 26-Nov-2021.) |
⊢ (Fun 𝑅 → ∅ ∉ 𝑅) | ||
Theorem | funss 5254 | Subclass theorem for function predicate. (Contributed by NM, 16-Aug-1994.) (Proof shortened by Mario Carneiro, 24-Jun-2014.) |
⊢ (𝐴 ⊆ 𝐵 → (Fun 𝐵 → Fun 𝐴)) | ||
Theorem | funeq 5255 | Equality theorem for function predicate. (Contributed by NM, 16-Aug-1994.) |
⊢ (𝐴 = 𝐵 → (Fun 𝐴 ↔ Fun 𝐵)) | ||
Theorem | funeqi 5256 | Equality inference for the function predicate. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ 𝐴 = 𝐵 ⇒ ⊢ (Fun 𝐴 ↔ Fun 𝐵) | ||
Theorem | funeqd 5257 | Equality deduction for the function predicate. (Contributed by NM, 23-Feb-2013.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (Fun 𝐴 ↔ Fun 𝐵)) | ||
Theorem | nffun 5258 | Bound-variable hypothesis builder for a function. (Contributed by NM, 30-Jan-2004.) |
⊢ Ⅎ𝑥𝐹 ⇒ ⊢ Ⅎ𝑥Fun 𝐹 | ||
Theorem | sbcfung 5259 | Distribute proper substitution through the function predicate. (Contributed by Alexander van der Vekens, 23-Jul-2017.) |
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]Fun 𝐹 ↔ Fun ⦋𝐴 / 𝑥⦌𝐹)) | ||
Theorem | funeu 5260* | There is exactly one value of a function. (Contributed by NM, 22-Apr-2004.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ ((Fun 𝐹 ∧ 𝐴𝐹𝐵) → ∃!𝑦 𝐴𝐹𝑦) | ||
Theorem | funeu2 5261* | There is exactly one value of a function. (Contributed by NM, 3-Aug-1994.) |
⊢ ((Fun 𝐹 ∧ 〈𝐴, 𝐵〉 ∈ 𝐹) → ∃!𝑦〈𝐴, 𝑦〉 ∈ 𝐹) | ||
Theorem | dffun7 5262* | Alternate definition of a function. One possibility for the definition of a function in [Enderton] p. 42. (Enderton's definition is ambiguous because "there is only one" could mean either "there is at most one" or "there is exactly one". However, dffun8 5263 shows that it does not matter which meaning we pick.) (Contributed by NM, 4-Nov-2002.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃*𝑦 𝑥𝐴𝑦)) | ||
Theorem | dffun8 5263* | Alternate definition of a function. One possibility for the definition of a function in [Enderton] p. 42. Compare dffun7 5262. (Contributed by NM, 4-Nov-2002.) (Proof shortened by Andrew Salmon, 17-Sep-2011.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃!𝑦 𝑥𝐴𝑦)) | ||
Theorem | dffun9 5264* | Alternate definition of a function. (Contributed by NM, 28-Mar-2007.) (Revised by NM, 16-Jun-2017.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃*𝑦 ∈ ran 𝐴 𝑥𝐴𝑦)) | ||
Theorem | funfn 5265 | An equivalence for the function predicate. (Contributed by NM, 13-Aug-2004.) |
⊢ (Fun 𝐴 ↔ 𝐴 Fn dom 𝐴) | ||
Theorem | funfnd 5266 | A function is a function over its domain. (Contributed by Glauco Siliprandi, 23-Oct-2021.) |
⊢ (𝜑 → Fun 𝐴) ⇒ ⊢ (𝜑 → 𝐴 Fn dom 𝐴) | ||
Theorem | funi 5267 | The identity relation is a function. Part of Theorem 10.4 of [Quine] p. 65. (Contributed by NM, 30-Apr-1998.) |
⊢ Fun I | ||
Theorem | nfunv 5268 | The universe is not a function. (Contributed by Raph Levien, 27-Jan-2004.) |
⊢ ¬ Fun V | ||
Theorem | funopg 5269 | A Kuratowski ordered pair is a function only if its components are equal. (Contributed by NM, 5-Jun-2008.) (Revised by Mario Carneiro, 26-Apr-2015.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊 ∧ Fun 〈𝐴, 𝐵〉) → 𝐴 = 𝐵) | ||
Theorem | funopab 5270* | 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 5271* | A class of ordered pairs of values is a function. (Contributed by NM, 14-Nov-1995.) |
⊢ Fun {〈𝑥, 𝑦〉 ∣ 𝑦 = 𝐴} | ||
Theorem | funopab4 5272* | A class of ordered pairs of values in the form used by df-mpt 4081 is a function. (Contributed by NM, 17-Feb-2013.) |
⊢ Fun {〈𝑥, 𝑦〉 ∣ (𝜑 ∧ 𝑦 = 𝐴)} | ||
Theorem | funmpt 5273 | A function in maps-to notation is a function. (Contributed by Mario Carneiro, 13-Jan-2013.) |
⊢ Fun (𝑥 ∈ 𝐴 ↦ 𝐵) | ||
Theorem | funmpt2 5274 | 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 5275 | 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 | funres 5276 | A restriction of a function is a function. Compare Exercise 18 of [TakeutiZaring] p. 25. (Contributed by NM, 16-Aug-1994.) |
⊢ (Fun 𝐹 → Fun (𝐹 ↾ 𝐴)) | ||
Theorem | funssres 5277 | The restriction of a function to the domain of a subclass equals the subclass. (Contributed by NM, 15-Aug-1994.) |
⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹) → (𝐹 ↾ dom 𝐺) = 𝐺) | ||
Theorem | fun2ssres 5278 | Equality of restrictions of a function and a subclass. (Contributed by NM, 16-Aug-1994.) |
⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹 ∧ 𝐴 ⊆ dom 𝐺) → (𝐹 ↾ 𝐴) = (𝐺 ↾ 𝐴)) | ||
Theorem | funun 5279 | 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 | funcnvsn 5280 | The converse singleton of an ordered pair is a function. This is equivalent to funsn 5283 via cnvsn 5129, but stating it this way allows us to skip the sethood assumptions on 𝐴 and 𝐵. (Contributed by NM, 30-Apr-2015.) |
⊢ Fun ◡{〈𝐴, 𝐵〉} | ||
Theorem | funsng 5281 | 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 5282 | Functionality and domain of the singleton of an ordered pair. (Contributed by Mario Carneiro, 30-Apr-2015.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → {〈𝐴, 𝐵〉} Fn {𝐴}) | ||
Theorem | funsn 5283 | 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 | funinsn 5284 | A function based on the singleton of an ordered pair. Unlike funsng 5281, this holds even if 𝐴 or 𝐵 is a proper class. (Contributed by Jim Kingdon, 17-Apr-2022.) |
⊢ Fun ({〈𝐴, 𝐵〉} ∩ (𝑉 × 𝑊)) | ||
Theorem | funprg 5285 | A set of two pairs is a function if their first members are different. (Contributed by FL, 26-Jun-2011.) |
⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) | ||
Theorem | funtpg 5286 | A set of three pairs is a function if their first members are different. (Contributed by Alexander van der Vekens, 5-Dec-2017.) |
⊢ (((𝑋 ∈ 𝑈 ∧ 𝑌 ∈ 𝑉 ∧ 𝑍 ∈ 𝑊) ∧ (𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐺 ∧ 𝐶 ∈ 𝐻) ∧ (𝑋 ≠ 𝑌 ∧ 𝑋 ≠ 𝑍 ∧ 𝑌 ≠ 𝑍)) → Fun {〈𝑋, 𝐴〉, 〈𝑌, 𝐵〉, 〈𝑍, 𝐶〉}) | ||
Theorem | funpr 5287 | A function with a domain of two elements. (Contributed by Jeff Madsen, 20-Jun-2010.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V ⇒ ⊢ (𝐴 ≠ 𝐵 → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) | ||
Theorem | funtp 5288 | A function with a domain of three elements. (Contributed by NM, 14-Sep-2011.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V & ⊢ 𝐸 ∈ V & ⊢ 𝐹 ∈ V ⇒ ⊢ ((𝐴 ≠ 𝐵 ∧ 𝐴 ≠ 𝐶 ∧ 𝐵 ≠ 𝐶) → Fun {〈𝐴, 𝐷〉, 〈𝐵, 𝐸〉, 〈𝐶, 𝐹〉}) | ||
Theorem | fnsn 5289 | Functionality and domain of the singleton of an ordered pair. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ {〈𝐴, 𝐵〉} Fn {𝐴} | ||
Theorem | fnprg 5290 | Function with a domain of two different values. (Contributed by FL, 26-Jun-2011.) (Revised by Mario Carneiro, 26-Apr-2015.) |
⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉} Fn {𝐴, 𝐵}) | ||
Theorem | fntpg 5291 | Function with a domain of three different values. (Contributed by Alexander van der Vekens, 5-Dec-2017.) |
⊢ (((𝑋 ∈ 𝑈 ∧ 𝑌 ∈ 𝑉 ∧ 𝑍 ∈ 𝑊) ∧ (𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐺 ∧ 𝐶 ∈ 𝐻) ∧ (𝑋 ≠ 𝑌 ∧ 𝑋 ≠ 𝑍 ∧ 𝑌 ≠ 𝑍)) → {〈𝑋, 𝐴〉, 〈𝑌, 𝐵〉, 〈𝑍, 𝐶〉} Fn {𝑋, 𝑌, 𝑍}) | ||
Theorem | fntp 5292 | 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 | fun0 5293 | The empty set is a function. Theorem 10.3 of [Quine] p. 65. (Contributed by NM, 7-Apr-1998.) |
⊢ Fun ∅ | ||
Theorem | funcnvcnv 5294 | The double converse of a function is a function. (Contributed by NM, 21-Sep-2004.) |
⊢ (Fun 𝐴 → Fun ◡◡𝐴) | ||
Theorem | funcnv2 5295* | A simpler equivalence for single-rooted (see funcnv 5296). (Contributed by NM, 9-Aug-2004.) |
⊢ (Fun ◡𝐴 ↔ ∀𝑦∃*𝑥 𝑥𝐴𝑦) | ||
Theorem | funcnv 5296* | 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 5295 for a simpler version. (Contributed by NM, 13-Aug-2004.) |
⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃*𝑥 𝑥𝐴𝑦) | ||
Theorem | funcnv3 5297* | A condition showing a class is single-rooted. (See funcnv 5296). (Contributed by NM, 26-May-2006.) |
⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃!𝑥 ∈ dom 𝐴 𝑥𝐴𝑦) | ||
Theorem | funcnveq 5298* | Another way of expressing that a class is single-rooted. Counterpart to dffun2 5245. (Contributed by Jim Kingdon, 24-Dec-2018.) |
⊢ (Fun ◡𝐴 ↔ ∀𝑥∀𝑦∀𝑧((𝑥𝐴𝑦 ∧ 𝑧𝐴𝑦) → 𝑥 = 𝑧)) | ||
Theorem | fun2cnv 5299* | 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 5300 | A single-valued relation is a function. (See fun2cnv 5299 for "single-valued.") Definition 6.4(4) of [TakeutiZaring] p. 24. (Contributed by NM, 17-Jan-2006.) |
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ Fun ◡◡𝐴)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |