Theorem List for Intuitionistic Logic Explorer - 5201-5300 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Definition | df-fn 5201 |
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 5202 |
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 5203 |
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 5204 |
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 5205 |
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 5206* |
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 4052), 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 5207* |
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 5208* |
Alternate definition of a function. (Contributed by NM,
29-Dec-1996.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((𝑥𝐴𝑦 ∧ 𝑥𝐴𝑧) → 𝑦 = 𝑧))) |
|
Theorem | dffun4 5209* |
Alternate definition of a function. Definition 6.4(4) of
[TakeutiZaring] p. 24.
(Contributed by NM, 29-Dec-1996.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((〈𝑥, 𝑦〉 ∈ 𝐴 ∧ 〈𝑥, 𝑧〉 ∈ 𝐴) → 𝑦 = 𝑧))) |
|
Theorem | dffun5r 5210* |
A way of proving a relation is a function, analogous to mo2r 2071.
(Contributed by Jim Kingdon, 27-May-2020.)
|
⊢ ((Rel 𝐴 ∧ ∀𝑥∃𝑧∀𝑦(〈𝑥, 𝑦〉 ∈ 𝐴 → 𝑦 = 𝑧)) → Fun 𝐴) |
|
Theorem | dffun6f 5211* |
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 5212* |
Alternate definition of a function using "at most one" notation.
(Contributed by NM, 9-Mar-1995.)
|
⊢ (Fun 𝐹 ↔ (Rel 𝐹 ∧ ∀𝑥∃*𝑦 𝑥𝐹𝑦)) |
|
Theorem | funmo 5213* |
A function has at most one value for each argument. (Contributed by NM,
24-May-1998.)
|
⊢ (Fun 𝐹 → ∃*𝑦 𝐴𝐹𝑦) |
|
Theorem | dffun4f 5214* |
Definition of function like dffun4 5209 but using bound-variable hypotheses
instead of distinct variable conditions. (Contributed by Jim Kingdon,
17-Mar-2019.)
|
⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑦𝐴
& ⊢ Ⅎ𝑧𝐴 ⇒ ⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥∀𝑦∀𝑧((〈𝑥, 𝑦〉 ∈ 𝐴 ∧ 〈𝑥, 𝑧〉 ∈ 𝐴) → 𝑦 = 𝑧))) |
|
Theorem | funrel 5215 |
A function is a relation. (Contributed by NM, 1-Aug-1994.)
|
⊢ (Fun 𝐴 → Rel 𝐴) |
|
Theorem | 0nelfun 5216 |
A function does not contain the empty set. (Contributed by BJ,
26-Nov-2021.)
|
⊢ (Fun 𝑅 → ∅ ∉ 𝑅) |
|
Theorem | funss 5217 |
Subclass theorem for function predicate. (Contributed by NM,
16-Aug-1994.) (Proof shortened by Mario Carneiro, 24-Jun-2014.)
|
⊢ (𝐴 ⊆ 𝐵 → (Fun 𝐵 → Fun 𝐴)) |
|
Theorem | funeq 5218 |
Equality theorem for function predicate. (Contributed by NM,
16-Aug-1994.)
|
⊢ (𝐴 = 𝐵 → (Fun 𝐴 ↔ Fun 𝐵)) |
|
Theorem | funeqi 5219 |
Equality inference for the function predicate. (Contributed by Jonathan
Ben-Naim, 3-Jun-2011.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (Fun 𝐴 ↔ Fun 𝐵) |
|
Theorem | funeqd 5220 |
Equality deduction for the function predicate. (Contributed by NM,
23-Feb-2013.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (Fun 𝐴 ↔ Fun 𝐵)) |
|
Theorem | nffun 5221 |
Bound-variable hypothesis builder for a function. (Contributed by NM,
30-Jan-2004.)
|
⊢ Ⅎ𝑥𝐹 ⇒ ⊢ Ⅎ𝑥Fun 𝐹 |
|
Theorem | sbcfung 5222 |
Distribute proper substitution through the function predicate.
(Contributed by Alexander van der Vekens, 23-Jul-2017.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]Fun 𝐹 ↔ Fun ⦋𝐴 / 𝑥⦌𝐹)) |
|
Theorem | funeu 5223* |
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 5224* |
There is exactly one value of a function. (Contributed by NM,
3-Aug-1994.)
|
⊢ ((Fun 𝐹 ∧ 〈𝐴, 𝐵〉 ∈ 𝐹) → ∃!𝑦〈𝐴, 𝑦〉 ∈ 𝐹) |
|
Theorem | dffun7 5225* |
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 5226 shows that it does not matter
which meaning we pick.) (Contributed by NM, 4-Nov-2002.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃*𝑦 𝑥𝐴𝑦)) |
|
Theorem | dffun8 5226* |
Alternate definition of a function. One possibility for the definition
of a function in [Enderton] p. 42.
Compare dffun7 5225. (Contributed by
NM, 4-Nov-2002.) (Proof shortened by Andrew Salmon, 17-Sep-2011.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃!𝑦 𝑥𝐴𝑦)) |
|
Theorem | dffun9 5227* |
Alternate definition of a function. (Contributed by NM, 28-Mar-2007.)
(Revised by NM, 16-Jun-2017.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∃*𝑦 ∈ ran 𝐴 𝑥𝐴𝑦)) |
|
Theorem | funfn 5228 |
An equivalence for the function predicate. (Contributed by NM,
13-Aug-2004.)
|
⊢ (Fun 𝐴 ↔ 𝐴 Fn dom 𝐴) |
|
Theorem | funfnd 5229 |
A function is a function over its domain. (Contributed by Glauco
Siliprandi, 23-Oct-2021.)
|
⊢ (𝜑 → Fun 𝐴) ⇒ ⊢ (𝜑 → 𝐴 Fn dom 𝐴) |
|
Theorem | funi 5230 |
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 5231 |
The universe is not a function. (Contributed by Raph Levien,
27-Jan-2004.)
|
⊢ ¬ Fun V |
|
Theorem | funopg 5232 |
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 5233* |
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 5234* |
A class of ordered pairs of values is a function. (Contributed by NM,
14-Nov-1995.)
|
⊢ Fun {〈𝑥, 𝑦〉 ∣ 𝑦 = 𝐴} |
|
Theorem | funopab4 5235* |
A class of ordered pairs of values in the form used by df-mpt 4052 is a
function. (Contributed by NM, 17-Feb-2013.)
|
⊢ Fun {〈𝑥, 𝑦〉 ∣ (𝜑 ∧ 𝑦 = 𝐴)} |
|
Theorem | funmpt 5236 |
A function in maps-to notation is a function. (Contributed by Mario
Carneiro, 13-Jan-2013.)
|
⊢ Fun (𝑥 ∈ 𝐴 ↦ 𝐵) |
|
Theorem | funmpt2 5237 |
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 5238 |
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 5239 |
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 5240 |
The restriction of a function to the domain of a subclass equals the
subclass. (Contributed by NM, 15-Aug-1994.)
|
⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹) → (𝐹 ↾ dom 𝐺) = 𝐺) |
|
Theorem | fun2ssres 5241 |
Equality of restrictions of a function and a subclass. (Contributed by
NM, 16-Aug-1994.)
|
⊢ ((Fun 𝐹 ∧ 𝐺 ⊆ 𝐹 ∧ 𝐴 ⊆ dom 𝐺) → (𝐹 ↾ 𝐴) = (𝐺 ↾ 𝐴)) |
|
Theorem | funun 5242 |
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 5243 |
The converse singleton of an ordered pair is a function. This is
equivalent to funsn 5246 via cnvsn 5093, but stating it this way allows us to
skip the sethood assumptions on 𝐴 and 𝐵. (Contributed by NM,
30-Apr-2015.)
|
⊢ Fun ◡{〈𝐴, 𝐵〉} |
|
Theorem | funsng 5244 |
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 5245 |
Functionality and domain of the singleton of an ordered pair.
(Contributed by Mario Carneiro, 30-Apr-2015.)
|
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → {〈𝐴, 𝐵〉} Fn {𝐴}) |
|
Theorem | funsn 5246 |
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 5247 |
A function based on the singleton of an ordered pair. Unlike funsng 5244,
this holds even if 𝐴 or 𝐵 is a proper class.
(Contributed by
Jim Kingdon, 17-Apr-2022.)
|
⊢ Fun ({〈𝐴, 𝐵〉} ∩ (𝑉 × 𝑊)) |
|
Theorem | funprg 5248 |
A set of two pairs is a function if their first members are different.
(Contributed by FL, 26-Jun-2011.)
|
⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) |
|
Theorem | funtpg 5249 |
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 5250 |
A function with a domain of two elements. (Contributed by Jeff Madsen,
20-Jun-2010.)
|
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈
V ⇒ ⊢ (𝐴 ≠ 𝐵 → Fun {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉}) |
|
Theorem | funtp 5251 |
A function with a domain of three elements. (Contributed by NM,
14-Sep-2011.)
|
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V & ⊢ 𝐸 ∈ V & ⊢ 𝐹 ∈
V ⇒ ⊢ ((𝐴 ≠ 𝐵 ∧ 𝐴 ≠ 𝐶 ∧ 𝐵 ≠ 𝐶) → Fun {〈𝐴, 𝐷〉, 〈𝐵, 𝐸〉, 〈𝐶, 𝐹〉}) |
|
Theorem | fnsn 5252 |
Functionality and domain of the singleton of an ordered pair.
(Contributed by Jonathan Ben-Naim, 3-Jun-2011.)
|
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈
V ⇒ ⊢ {〈𝐴, 𝐵〉} Fn {𝐴} |
|
Theorem | fnprg 5253 |
Function with a domain of two different values. (Contributed by FL,
26-Jun-2011.) (Revised by Mario Carneiro, 26-Apr-2015.)
|
⊢ (((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) ∧ (𝐶 ∈ 𝑋 ∧ 𝐷 ∈ 𝑌) ∧ 𝐴 ≠ 𝐵) → {〈𝐴, 𝐶〉, 〈𝐵, 𝐷〉} Fn {𝐴, 𝐵}) |
|
Theorem | fntpg 5254 |
Function with a domain of three different values. (Contributed by
Alexander van der Vekens, 5-Dec-2017.)
|
⊢ (((𝑋 ∈ 𝑈 ∧ 𝑌 ∈ 𝑉 ∧ 𝑍 ∈ 𝑊) ∧ (𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐺 ∧ 𝐶 ∈ 𝐻) ∧ (𝑋 ≠ 𝑌 ∧ 𝑋 ≠ 𝑍 ∧ 𝑌 ≠ 𝑍)) → {〈𝑋, 𝐴〉, 〈𝑌, 𝐵〉, 〈𝑍, 𝐶〉} Fn {𝑋, 𝑌, 𝑍}) |
|
Theorem | fntp 5255 |
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 5256 |
The empty set is a function. Theorem 10.3 of [Quine] p. 65. (Contributed
by NM, 7-Apr-1998.)
|
⊢ Fun ∅ |
|
Theorem | funcnvcnv 5257 |
The double converse of a function is a function. (Contributed by NM,
21-Sep-2004.)
|
⊢ (Fun 𝐴 → Fun ◡◡𝐴) |
|
Theorem | funcnv2 5258* |
A simpler equivalence for single-rooted (see funcnv 5259). (Contributed
by NM, 9-Aug-2004.)
|
⊢ (Fun ◡𝐴 ↔ ∀𝑦∃*𝑥 𝑥𝐴𝑦) |
|
Theorem | funcnv 5259* |
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 5258 for a simpler version. (Contributed by
NM, 13-Aug-2004.)
|
⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃*𝑥 𝑥𝐴𝑦) |
|
Theorem | funcnv3 5260* |
A condition showing a class is single-rooted. (See funcnv 5259).
(Contributed by NM, 26-May-2006.)
|
⊢ (Fun ◡𝐴 ↔ ∀𝑦 ∈ ran 𝐴∃!𝑥 ∈ dom 𝐴 𝑥𝐴𝑦) |
|
Theorem | funcnveq 5261* |
Another way of expressing that a class is single-rooted. Counterpart to
dffun2 5208. (Contributed by Jim Kingdon, 24-Dec-2018.)
|
⊢ (Fun ◡𝐴 ↔ ∀𝑥∀𝑦∀𝑧((𝑥𝐴𝑦 ∧ 𝑧𝐴𝑦) → 𝑥 = 𝑧)) |
|
Theorem | fun2cnv 5262* |
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 5263 |
A single-valued relation is a function. (See fun2cnv 5262 for
"single-valued.") Definition 6.4(4) of [TakeutiZaring] p. 24.
(Contributed by NM, 17-Jan-2006.)
|
⊢ (Fun 𝐴 ↔ (Rel 𝐴 ∧ Fun ◡◡𝐴)) |
|
Theorem | fncnv 5264* |
Single-rootedness (see funcnv 5259) of a class cut down by a cross
product. (Contributed by NM, 5-Mar-2007.)
|
⊢ (◡(𝑅 ∩ (𝐴 × 𝐵)) Fn 𝐵 ↔ ∀𝑦 ∈ 𝐵 ∃!𝑥 ∈ 𝐴 𝑥𝑅𝑦) |
|
Theorem | fun11 5265* |
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 5266* |
The union of a chain (with respect to inclusion) of functions is a
function. (Contributed by NM, 10-Aug-2004.)
|
⊢ (∀𝑓 ∈ 𝐴 (Fun 𝑓 ∧ ∀𝑔 ∈ 𝐴 (𝑓 ⊆ 𝑔 ∨ 𝑔 ⊆ 𝑓)) → Fun ∪
𝐴) |
|
Theorem | funcnvuni 5267* |
The union of a chain (with respect to inclusion) of single-rooted sets
is single-rooted. (See funcnv 5259 for "single-rooted"
definition.)
(Contributed by NM, 11-Aug-2004.)
|
⊢ (∀𝑓 ∈ 𝐴 (Fun ◡𝑓 ∧ ∀𝑔 ∈ 𝐴 (𝑓 ⊆ 𝑔 ∨ 𝑔 ⊆ 𝑓)) → Fun ◡∪ 𝐴) |
|
Theorem | fun11uni 5268* |
The union of a chain (with respect to inclusion) of one-to-one functions
is a one-to-one function. (Contributed by NM, 11-Aug-2004.)
|
⊢ (∀𝑓 ∈ 𝐴 ((Fun 𝑓 ∧ Fun ◡𝑓) ∧ ∀𝑔 ∈ 𝐴 (𝑓 ⊆ 𝑔 ∨ 𝑔 ⊆ 𝑓)) → (Fun ∪
𝐴 ∧ Fun ◡∪ 𝐴)) |
|
Theorem | funin 5269 |
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 5270 |
The restriction of a one-to-one function is one-to-one. (Contributed by
NM, 25-Mar-1998.)
|
⊢ (Fun ◡𝐹 → Fun ◡(𝐹 ↾ 𝐴)) |
|
Theorem | funcnvres 5271 |
The converse of a restricted function. (Contributed by NM,
27-Mar-1998.)
|
⊢ (Fun ◡𝐹 → ◡(𝐹 ↾ 𝐴) = (◡𝐹 ↾ (𝐹 “ 𝐴))) |
|
Theorem | cnvresid 5272 |
Converse of a restricted identity function. (Contributed by FL,
4-Mar-2007.)
|
⊢ ◡( I
↾ 𝐴) = ( I ↾
𝐴) |
|
Theorem | funcnvres2 5273 |
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 5274 |
The image of the preimage of a function. (Contributed by NM,
25-May-2004.)
|
⊢ (Fun 𝐹 → (𝐹 “ (◡𝐹 “ 𝐴)) = (𝐴 ∩ ran 𝐹)) |
|
Theorem | funimass1 5275 |
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 5276 |
A kind of contraposition law that infers an image subclass from a subclass
of a preimage. (Contributed by NM, 25-May-2004.)
|
⊢ ((Fun 𝐹 ∧ 𝐴 ⊆ (◡𝐹 “ 𝐵)) → (𝐹 “ 𝐴) ⊆ 𝐵) |
|
Theorem | imadiflem 5277 |
One direction of imadif 5278. This direction does not require
Fun ◡𝐹. (Contributed by Jim
Kingdon, 25-Dec-2018.)
|
⊢ ((𝐹 “ 𝐴) ∖ (𝐹 “ 𝐵)) ⊆ (𝐹 “ (𝐴 ∖ 𝐵)) |
|
Theorem | imadif 5278 |
The image of a difference is the difference of images. (Contributed by
NM, 24-May-1998.)
|
⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∖ 𝐵)) = ((𝐹 “ 𝐴) ∖ (𝐹 “ 𝐵))) |
|
Theorem | imainlem 5279 |
One direction of imain 5280. This direction does not require
Fun ◡𝐹. (Contributed by Jim
Kingdon, 25-Dec-2018.)
|
⊢ (𝐹 “ (𝐴 ∩ 𝐵)) ⊆ ((𝐹 “ 𝐴) ∩ (𝐹 “ 𝐵)) |
|
Theorem | imain 5280 |
The image of an intersection is the intersection of images.
(Contributed by Paul Chapman, 11-Apr-2009.)
|
⊢ (Fun ◡𝐹 → (𝐹 “ (𝐴 ∩ 𝐵)) = ((𝐹 “ 𝐴) ∩ (𝐹 “ 𝐵))) |
|
Theorem | funimaexglem 5281 |
Lemma for funimaexg 5282. It constitutes the interesting part of
funimaexg 5282, in which 𝐵 ⊆ dom 𝐴. (Contributed by Jim Kingdon,
27-Dec-2018.)
|
⊢ ((Fun 𝐴 ∧ 𝐵 ∈ 𝐶 ∧ 𝐵 ⊆ dom 𝐴) → (𝐴 “ 𝐵) ∈ V) |
|
Theorem | funimaexg 5282 |
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 5283 |
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) |
|
Theorem | isarep1 5284* |
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 5285* |
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 5283. (Contributed by NM, 26-Oct-2006.)
|
⊢ 𝐴 ∈ V & ⊢ ∀𝑥 ∈ 𝐴 ∀𝑦∀𝑧((𝜑 ∧ [𝑧 / 𝑦]𝜑) → 𝑦 = 𝑧) ⇒ ⊢ ∃𝑤 𝑤 = ({〈𝑥, 𝑦〉 ∣ 𝜑} “ 𝐴) |
|
Theorem | fneq1 5286 |
Equality theorem for function predicate with domain. (Contributed by NM,
1-Aug-1994.)
|
⊢ (𝐹 = 𝐺 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) |
|
Theorem | fneq2 5287 |
Equality theorem for function predicate with domain. (Contributed by NM,
1-Aug-1994.)
|
⊢ (𝐴 = 𝐵 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) |
|
Theorem | fneq1d 5288 |
Equality deduction for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
⊢ (𝜑 → 𝐹 = 𝐺) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴)) |
|
Theorem | fneq2d 5289 |
Equality deduction for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵)) |
|
Theorem | fneq12d 5290 |
Equality deduction for function predicate with domain. (Contributed by
NM, 26-Jun-2011.)
|
⊢ (𝜑 → 𝐹 = 𝐺)
& ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) |
|
Theorem | fneq12 5291 |
Equality theorem for function predicate with domain. (Contributed by
Thierry Arnoux, 31-Jan-2017.)
|
⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐵) → (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐵)) |
|
Theorem | fneq1i 5292 |
Equality inference for function predicate with domain. (Contributed by
Paul Chapman, 22-Jun-2011.)
|
⊢ 𝐹 = 𝐺 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐺 Fn 𝐴) |
|
Theorem | fneq2i 5293 |
Equality inference for function predicate with domain. (Contributed by
NM, 4-Sep-2011.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐹 Fn 𝐴 ↔ 𝐹 Fn 𝐵) |
|
Theorem | nffn 5294 |
Bound-variable hypothesis builder for a function with domain.
(Contributed by NM, 30-Jan-2004.)
|
⊢ Ⅎ𝑥𝐹
& ⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥 𝐹 Fn 𝐴 |
|
Theorem | fnfun 5295 |
A function with domain is a function. (Contributed by NM, 1-Aug-1994.)
|
⊢ (𝐹 Fn 𝐴 → Fun 𝐹) |
|
Theorem | fnrel 5296 |
A function with domain is a relation. (Contributed by NM, 1-Aug-1994.)
|
⊢ (𝐹 Fn 𝐴 → Rel 𝐹) |
|
Theorem | fndm 5297 |
The domain of a function. (Contributed by NM, 2-Aug-1994.)
|
⊢ (𝐹 Fn 𝐴 → dom 𝐹 = 𝐴) |
|
Theorem | funfni 5298 |
Inference to convert a function and domain antecedent. (Contributed by
NM, 22-Apr-2004.)
|
⊢ ((Fun 𝐹 ∧ 𝐵 ∈ dom 𝐹) → 𝜑) ⇒ ⊢ ((𝐹 Fn 𝐴 ∧ 𝐵 ∈ 𝐴) → 𝜑) |
|
Theorem | fndmu 5299 |
A function has a unique domain. (Contributed by NM, 11-Aug-1994.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐹 Fn 𝐵) → 𝐴 = 𝐵) |
|
Theorem | fnbr 5300 |
The first argument of binary relation on a function belongs to the
function's domain. (Contributed by NM, 7-May-2004.)
|
⊢ ((𝐹 Fn 𝐴 ∧ 𝐵𝐹𝐶) → 𝐵 ∈ 𝐴) |