![]() |
Metamath
Proof Explorer Theorem List (p. 86 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 | card2inf 8501* | The definition cardval2 8855 has the curious property that for non-numerable sets (for which ndmfv 6256 yields ∅), it still evaluates to a nonempty set, and indeed it contains ω. (Contributed by Mario Carneiro, 15-Jan-2013.) (Revised by Mario Carneiro, 27-Apr-2015.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (¬ ∃𝑦 ∈ On 𝑦 ≈ 𝐴 → ω ⊆ {𝑥 ∈ On ∣ 𝑥 ≺ 𝐴}) | ||
Syntax | char 8502 | Class symbol for the Hartogs/cardinal successor function. |
class har | ||
Syntax | cwdom 8503 | Class symbol for the weak dominance relation. |
class ≼* | ||
Definition | df-har 8504* |
Define the Hartogs function , which maps all sets to the smallest
ordinal that cannot be injected into the given set. In the important
special case where 𝑥 is an ordinal, this is the
cardinal successor
operation.
Traditionally, the Hartogs number of a set is written ℵ(𝑋) and the cardinal successor 𝑋 +; we use functional notation for this, and cannot use the aleph symbol because it is taken for the enumerating function of the infinite initial ordinals df-aleph 8804. Some authors define the Hartogs number of a set to be the least *infinite* ordinal which does not inject into it, thus causing the range to consist only of alephs. We use the simpler definition where the value can be any successor cardinal. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ har = (𝑥 ∈ V ↦ {𝑦 ∈ On ∣ 𝑦 ≼ 𝑥}) | ||
Definition | df-wdom 8505* | A set is weakly dominated by a "larger" set iff the "larger" set can be mapped onto the "smaller" set or the smaller set is empty; equivalently if the smaller set can be placed into bijection with some partition of the larger set. When choice is assumed (as fodom 9382), this coincides with the 1-1 definition df-dom 7999; however, it is not known whether this is a choice-equivalent or a strictly weaker form. Some discussion of this question can be found at http://boolesrings.org/asafk/2014/on-the-partition-principle/. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ ≼* = {〈𝑥, 𝑦〉 ∣ (𝑥 = ∅ ∨ ∃𝑧 𝑧:𝑦–onto→𝑥)} | ||
Theorem | harf 8506 | Functionality of the Hartogs function. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ har:V⟶On | ||
Theorem | harcl 8507 | Closure of the Hartogs function in the ordinals. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (har‘𝑋) ∈ On | ||
Theorem | harval 8508* | Function value of the Hartogs function. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑋 ∈ 𝑉 → (har‘𝑋) = {𝑦 ∈ On ∣ 𝑦 ≼ 𝑋}) | ||
Theorem | elharval 8509 | The Hartogs number of a set is greater than all ordinals which inject into it. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 15-May-2015.) |
⊢ (𝑌 ∈ (har‘𝑋) ↔ (𝑌 ∈ On ∧ 𝑌 ≼ 𝑋)) | ||
Theorem | harndom 8510 | The Hartogs number of a set does not inject into that set. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 15-May-2015.) |
⊢ ¬ (har‘𝑋) ≼ 𝑋 | ||
Theorem | harword 8511 | Weak ordering property of the Hartogs function. (Contributed by Stefan O'Rear, 14-Feb-2015.) |
⊢ (𝑋 ≼ 𝑌 → (har‘𝑋) ⊆ (har‘𝑌)) | ||
Theorem | relwdom 8512 | Weak dominance is a relation. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ Rel ≼* | ||
Theorem | brwdom 8513* | Property of weak dominance (definitional form). (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑌 ∈ 𝑉 → (𝑋 ≼* 𝑌 ↔ (𝑋 = ∅ ∨ ∃𝑧 𝑧:𝑌–onto→𝑋))) | ||
Theorem | brwdomi 8514* | Property of weak dominance, forward direction only. (Contributed by Mario Carneiro, 5-May-2015.) |
⊢ (𝑋 ≼* 𝑌 → (𝑋 = ∅ ∨ ∃𝑧 𝑧:𝑌–onto→𝑋)) | ||
Theorem | brwdomn0 8515* | Weak dominance over nonempty sets. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 5-May-2015.) |
⊢ (𝑋 ≠ ∅ → (𝑋 ≼* 𝑌 ↔ ∃𝑧 𝑧:𝑌–onto→𝑋)) | ||
Theorem | 0wdom 8516 | Any set weakly dominates the empty set. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑋 ∈ 𝑉 → ∅ ≼* 𝑋) | ||
Theorem | fowdom 8517 | An onto function implies weak dominance. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ ((𝐹 ∈ 𝑉 ∧ 𝐹:𝑌–onto→𝑋) → 𝑋 ≼* 𝑌) | ||
Theorem | wdomref 8518 | Reflexivity of weak dominance. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑋 ∈ 𝑉 → 𝑋 ≼* 𝑋) | ||
Theorem | brwdom2 8519* | Alternate characterization of the weak dominance predicate which does not require special treatment of the empty set. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑌 ∈ 𝑉 → (𝑋 ≼* 𝑌 ↔ ∃𝑦 ∈ 𝒫 𝑌∃𝑧 𝑧:𝑦–onto→𝑋)) | ||
Theorem | domwdom 8520 | Weak dominance is implied by dominance in the usual sense. (Contributed by Stefan O'Rear, 11-Feb-2015.) |
⊢ (𝑋 ≼ 𝑌 → 𝑋 ≼* 𝑌) | ||
Theorem | wdomtr 8521 | Transitivity of weak dominance. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 5-May-2015.) |
⊢ ((𝑋 ≼* 𝑌 ∧ 𝑌 ≼* 𝑍) → 𝑋 ≼* 𝑍) | ||
Theorem | wdomen1 8522 | Equality-like theorem for equinumerosity and weak dominance. (Contributed by Mario Carneiro, 18-May-2015.) |
⊢ (𝐴 ≈ 𝐵 → (𝐴 ≼* 𝐶 ↔ 𝐵 ≼* 𝐶)) | ||
Theorem | wdomen2 8523 | Equality-like theorem for equinumerosity and weak dominance. (Contributed by Mario Carneiro, 18-May-2015.) |
⊢ (𝐴 ≈ 𝐵 → (𝐶 ≼* 𝐴 ↔ 𝐶 ≼* 𝐵)) | ||
Theorem | wdompwdom 8524 | Weak dominance strengthens to usual dominance on the power sets. (Contributed by Stefan O'Rear, 11-Feb-2015.) (Revised by Mario Carneiro, 5-May-2015.) |
⊢ (𝑋 ≼* 𝑌 → 𝒫 𝑋 ≼ 𝒫 𝑌) | ||
Theorem | canthwdom 8525 | Cantor's Theorem, stated using weak dominance (this is actually a stronger statement than canth2 8154, equivalent to canth 6648). (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ ¬ 𝒫 𝐴 ≼* 𝐴 | ||
Theorem | wdom2d 8526* | Deduce weak dominance from an implicit onto function (stated in a way which avoids ax-rep 4804). (Contributed by Stefan O'Rear, 13-Feb-2015.) |
⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐵 ∈ 𝑊) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∃𝑦 ∈ 𝐵 𝑥 = 𝑋) ⇒ ⊢ (𝜑 → 𝐴 ≼* 𝐵) | ||
Theorem | wdomd 8527* | Deduce weak dominance from an implicit onto function. (Contributed by Stefan O'Rear, 13-Feb-2015.) |
⊢ (𝜑 → 𝐵 ∈ 𝑊) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∃𝑦 ∈ 𝐵 𝑥 = 𝑋) ⇒ ⊢ (𝜑 → 𝐴 ≼* 𝐵) | ||
Theorem | brwdom3 8528* | Condition for weak dominance with a condition reminiscent of wdomd 8527. (Contributed by Stefan O'Rear, 13-Feb-2015.) (Revised by Mario Carneiro, 25-Jun-2015.) |
⊢ ((𝑋 ∈ 𝑉 ∧ 𝑌 ∈ 𝑊) → (𝑋 ≼* 𝑌 ↔ ∃𝑓∀𝑥 ∈ 𝑋 ∃𝑦 ∈ 𝑌 𝑥 = (𝑓‘𝑦))) | ||
Theorem | brwdom3i 8529* | Weak dominance implies existence of a covering function. (Contributed by Stefan O'Rear, 13-Feb-2015.) |
⊢ (𝑋 ≼* 𝑌 → ∃𝑓∀𝑥 ∈ 𝑋 ∃𝑦 ∈ 𝑌 𝑥 = (𝑓‘𝑦)) | ||
Theorem | unwdomg 8530 | Weak dominance of a (disjoint) union. (Contributed by Stefan O'Rear, 13-Feb-2015.) (Revised by Mario Carneiro, 25-Jun-2015.) |
⊢ ((𝐴 ≼* 𝐵 ∧ 𝐶 ≼* 𝐷 ∧ (𝐵 ∩ 𝐷) = ∅) → (𝐴 ∪ 𝐶) ≼* (𝐵 ∪ 𝐷)) | ||
Theorem | xpwdomg 8531 | Weak dominance of a Cartesian product. (Contributed by Stefan O'Rear, 13-Feb-2015.) (Revised by Mario Carneiro, 25-Jun-2015.) |
⊢ ((𝐴 ≼* 𝐵 ∧ 𝐶 ≼* 𝐷) → (𝐴 × 𝐶) ≼* (𝐵 × 𝐷)) | ||
Theorem | wdomima2g 8532 | A set is weakly dominant over its image under any function. This version of wdomimag 8533 is stated so as to avoid ax-rep 4804. (Contributed by Mario Carneiro, 25-Jun-2015.) |
⊢ ((Fun 𝐹 ∧ 𝐴 ∈ 𝑉 ∧ (𝐹 “ 𝐴) ∈ 𝑊) → (𝐹 “ 𝐴) ≼* 𝐴) | ||
Theorem | wdomimag 8533 | A set is weakly dominant over its image under any function. (Contributed by Stefan O'Rear, 14-Feb-2015.) (Revised by Mario Carneiro, 25-Jun-2015.) |
⊢ ((Fun 𝐹 ∧ 𝐴 ∈ 𝑉) → (𝐹 “ 𝐴) ≼* 𝐴) | ||
Theorem | unxpwdom2 8534 | Lemma for unxpwdom 8535. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ ((𝐴 × 𝐴) ≈ (𝐵 ∪ 𝐶) → (𝐴 ≼* 𝐵 ∨ 𝐴 ≼ 𝐶)) | ||
Theorem | unxpwdom 8535 | If a Cartesian product is dominated by a union, then the base set is either weakly dominated by one factor of the union or dominated by the other. Extracted from Lemma 2.3 of [KanamoriPincus] p. 420. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ ((𝐴 × 𝐴) ≼ (𝐵 ∪ 𝐶) → (𝐴 ≼* 𝐵 ∨ 𝐴 ≼ 𝐶)) | ||
Theorem | harwdom 8536 | The Hartogs function is weakly dominated by 𝒫 (𝑋 × 𝑋). This follows from a more precise analysis of the bound used in hartogs 8490 to prove that (har‘𝑋) is a set. (Contributed by Mario Carneiro, 15-May-2015.) |
⊢ (𝑋 ∈ 𝑉 → (har‘𝑋) ≼* 𝒫 (𝑋 × 𝑋)) | ||
Theorem | ixpiunwdom 8537* | Describe an onto function from the indexed cartesian product to the indexed union. Together with ixpssmapg 7980 this shows that ∪ 𝑥 ∈ 𝐴𝐵 and X𝑥 ∈ 𝐴𝐵 have closely linked cardinalities. (Contributed by Mario Carneiro, 27-Aug-2015.) |
⊢ ((𝐴 ∈ 𝑉 ∧ ∪ 𝑥 ∈ 𝐴 𝐵 ∈ 𝑊 ∧ X𝑥 ∈ 𝐴 𝐵 ≠ ∅) → ∪ 𝑥 ∈ 𝐴 𝐵 ≼* (X𝑥 ∈ 𝐴 𝐵 × 𝐴)) | ||
Axiom | ax-reg 8538* | Axiom of Regularity. An axiom of Zermelo-Fraenkel set theory. Also called the Axiom of Foundation. A rather non-intuitive axiom that denies more than it asserts, it states (in the form of zfreg 8541) that every nonempty set contains a set disjoint from itself. One consequence is that it denies the existence of a set containing itself (elirrv 8542). A stronger version that works for proper classes is proved as zfregs 8646. (Contributed by NM, 14-Aug-1993.) |
⊢ (∃𝑦 𝑦 ∈ 𝑥 → ∃𝑦(𝑦 ∈ 𝑥 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ¬ 𝑧 ∈ 𝑥))) | ||
Theorem | axreg2 8539* | Axiom of Regularity expressed more compactly. (Contributed by NM, 14-Aug-2003.) |
⊢ (𝑥 ∈ 𝑦 → ∃𝑥(𝑥 ∈ 𝑦 ∧ ∀𝑧(𝑧 ∈ 𝑥 → ¬ 𝑧 ∈ 𝑦))) | ||
Theorem | zfregcl 8540* | The Axiom of Regularity with class variables. (Contributed by NM, 5-Aug-1994.) Replace sethood hypothesis with sethood antecedent. (Revised by BJ, 27-Apr-2021.) |
⊢ (𝐴 ∈ 𝑉 → (∃𝑥 𝑥 ∈ 𝐴 → ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝑥 ¬ 𝑦 ∈ 𝐴)) | ||
Theorem | zfreg 8541* | The Axiom of Regularity using abbreviations. Axiom 6 of [TakeutiZaring] p. 21. This is called the "weak form." Axiom Reg of [BellMachover] p. 480. There is also a "strong form," not requiring that 𝐴 be a set, that can be proved with more difficulty (see zfregs 8646). (Contributed by NM, 26-Nov-1995.) Replace sethood hypothesis with sethood antecedent. (Revised by BJ, 27-Apr-2021.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐴 ≠ ∅) → ∃𝑥 ∈ 𝐴 (𝑥 ∩ 𝐴) = ∅) | ||
Theorem | elirrv 8542 | The membership relation is irreflexive: no set is a member of itself. Theorem 105 of [Suppes] p. 54. (This is trivial to prove from zfregfr 8547 and efrirr 5124, but this proof is direct from the Axiom of Regularity.) (Contributed by NM, 19-Aug-1993.) |
⊢ ¬ 𝑥 ∈ 𝑥 | ||
Theorem | elirr 8543 | No class is a member of itself. Exercise 6 of [TakeutiZaring] p. 22. (Contributed by NM, 7-Aug-1994.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ ¬ 𝐴 ∈ 𝐴 | ||
Theorem | sucprcreg 8544 | A class is equal to its successor iff it is a proper class (assuming the Axiom of Regularity). (Contributed by NM, 9-Jul-2004.) (Proof shortened by BJ, 16-Apr-2019.) |
⊢ (¬ 𝐴 ∈ V ↔ suc 𝐴 = 𝐴) | ||
Theorem | ruv 8545 | The Russell class is equal to the universe V. Exercise 5 of [TakeutiZaring] p. 22. (Contributed by Alan Sare, 4-Oct-2008.) |
⊢ {𝑥 ∣ 𝑥 ∉ 𝑥} = V | ||
Theorem | ruALT 8546 | Alternate proof of ru 3467, simplified using (indirectly) the Axiom of Regularity ax-reg 8538. (Contributed by Alan Sare, 4-Oct-2008.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ {𝑥 ∣ 𝑥 ∉ 𝑥} ∉ V | ||
Theorem | zfregfr 8547 | The epsilon relation is well-founded on any class. (Contributed by NM, 26-Nov-1995.) |
⊢ E Fr 𝐴 | ||
Theorem | en2lp 8548 | No class has 2-cycle membership loops. Theorem 7X(b) of [Enderton] p. 206. (Contributed by NM, 16-Oct-1996.) (Revised by Mario Carneiro, 25-Jun-2015.) |
⊢ ¬ (𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐴) | ||
Theorem | en3lplem1 8549* | Lemma for en3lp 8551. (Contributed by Alan Sare, 28-Oct-2011.) |
⊢ ((𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐶 ∧ 𝐶 ∈ 𝐴) → (𝑥 = 𝐴 → (𝑥 ∩ {𝐴, 𝐵, 𝐶}) ≠ ∅)) | ||
Theorem | en3lplem2 8550* | Lemma for en3lp 8551. (Contributed by Alan Sare, 28-Oct-2011.) |
⊢ ((𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐶 ∧ 𝐶 ∈ 𝐴) → (𝑥 ∈ {𝐴, 𝐵, 𝐶} → (𝑥 ∩ {𝐴, 𝐵, 𝐶}) ≠ ∅)) | ||
Theorem | en3lp 8551 | No class has 3-cycle membership loops. This proof was automatically generated from the virtual deduction proof en3lpVD 39394 using a translation program. (Contributed by Alan Sare, 24-Oct-2011.) |
⊢ ¬ (𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐶 ∧ 𝐶 ∈ 𝐴) | ||
Theorem | preleq 8552 | Equality of two unordered pairs when one member of each pair contains the other member. (Contributed by NM, 16-Oct-1996.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V ⇒ ⊢ (((𝐴 ∈ 𝐵 ∧ 𝐶 ∈ 𝐷) ∧ {𝐴, 𝐵} = {𝐶, 𝐷}) → (𝐴 = 𝐶 ∧ 𝐵 = 𝐷)) | ||
Theorem | opthreg 8553 | Theorem for alternate representation of ordered pairs, requiring the Axiom of Regularity ax-reg 8538 (via the preleq 8552 step). See df-op 4217 for a description of other ordered pair representations. Exercise 34 of [Enderton] p. 207. (Contributed by NM, 16-Oct-1996.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ 𝐷 ∈ V ⇒ ⊢ ({𝐴, {𝐴, 𝐵}} = {𝐶, {𝐶, 𝐷}} ↔ (𝐴 = 𝐶 ∧ 𝐵 = 𝐷)) | ||
Theorem | suc11reg 8554 | The successor operation behaves like a one-to-one function (assuming the Axiom of Regularity). Exercise 35 of [Enderton] p. 208 and its converse. (Contributed by NM, 25-Oct-2003.) |
⊢ (suc 𝐴 = suc 𝐵 ↔ 𝐴 = 𝐵) | ||
Theorem | dford2 8555* | Assuming ax-reg 8538, an ordinal is a transitive class on which inclusion satisfies trichotomy. (Contributed by Scott Fenton, 27-Oct-2010.) |
⊢ (Ord 𝐴 ↔ (Tr 𝐴 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥))) | ||
Theorem | inf0 8556* | Our Axiom of Infinity derived from existence of omega. The proof shows that the especially contrived class "ran (rec((𝑣 ∈ V ↦ suc 𝑣), 𝑥) ↾ ω) " exists, is a subset of its union, and contains a given set 𝑥 (and thus is nonempty). Thus, it provides an example demonstrating that a set 𝑦 exists with the necessary properties demanded by ax-inf 8573. (Contributed by NM, 15-Oct-1996.) |
⊢ ω ∈ V ⇒ ⊢ ∃𝑦(𝑥 ∈ 𝑦 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ∃𝑤(𝑧 ∈ 𝑤 ∧ 𝑤 ∈ 𝑦))) | ||
Theorem | inf1 8557 | Variation of Axiom of Infinity (using zfinf 8574 as a hypothesis). Axiom of Infinity in [FreydScedrov] p. 283. (Contributed by NM, 14-Oct-1996.) (Revised by David Abernethy, 1-Oct-2013.) |
⊢ ∃𝑥(𝑦 ∈ 𝑥 ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑥))) ⇒ ⊢ ∃𝑥(𝑥 ≠ ∅ ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑥))) | ||
Theorem | inf2 8558* | Variation of Axiom of Infinity. There exists a nonempty set that is a subset of its union (using zfinf 8574 as a hypothesis). Abbreviated version of the Axiom of Infinity in [FreydScedrov] p. 283. (Contributed by NM, 28-Oct-1996.) |
⊢ ∃𝑥(𝑦 ∈ 𝑥 ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑥))) ⇒ ⊢ ∃𝑥(𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) | ||
Theorem | inf3lema 8559* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ∈ (𝐺‘𝐵) ↔ (𝐴 ∈ 𝑥 ∧ (𝐴 ∩ 𝑥) ⊆ 𝐵)) | ||
Theorem | inf3lemb 8560* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐹‘∅) = ∅ | ||
Theorem | inf3lemc 8561* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ∈ ω → (𝐹‘suc 𝐴) = (𝐺‘(𝐹‘𝐴))) | ||
Theorem | inf3lemd 8562* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ∈ ω → (𝐹‘𝐴) ⊆ 𝑥) | ||
Theorem | inf3lem1 8563* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ (𝐴 ∈ ω → (𝐹‘𝐴) ⊆ (𝐹‘suc 𝐴)) | ||
Theorem | inf3lem2 8564* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 28-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → (𝐴 ∈ ω → (𝐹‘𝐴) ≠ 𝑥)) | ||
Theorem | inf3lem3 8565* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. In the proof, we invoke the Axiom of Regularity in the form of zfreg 8541. (Contributed by NM, 29-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → (𝐴 ∈ ω → (𝐹‘𝐴) ≠ (𝐹‘suc 𝐴))) | ||
Theorem | inf3lem4 8566* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 29-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → (𝐴 ∈ ω → (𝐹‘𝐴) ⊊ (𝐹‘suc 𝐴))) | ||
Theorem | inf3lem5 8567* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 29-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → ((𝐴 ∈ ω ∧ 𝐵 ∈ 𝐴) → (𝐹‘𝐵) ⊊ (𝐹‘𝐴))) | ||
Theorem | inf3lem6 8568* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. (Contributed by NM, 29-Oct-1996.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → 𝐹:ω–1-1→𝒫 𝑥) | ||
Theorem | inf3lem7 8569* | Lemma for our Axiom of Infinity => standard Axiom of Infinity. See inf3 8570 for detailed description. In the proof, we invoke the Axiom of Replacement in the form of f1dmex 7178. (Contributed by NM, 29-Oct-1996.) (Proof shortened by Mario Carneiro, 19-Jan-2013.) |
⊢ 𝐺 = (𝑦 ∈ V ↦ {𝑤 ∈ 𝑥 ∣ (𝑤 ∩ 𝑥) ⊆ 𝑦}) & ⊢ 𝐹 = (rec(𝐺, ∅) ↾ ω) & ⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ((𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) → ω ∈ V) | ||
Theorem | inf3 8570 |
Our Axiom of Infinity ax-inf 8573 implies the standard Axiom of Infinity.
The hypothesis is a variant of our Axiom of Infinity provided by
inf2 8558, and the conclusion is the version of the Axiom of Infinity
shown as Axiom 7 in [TakeutiZaring] p. 43. (Other standard versions are
proved later as axinf2 8575 and zfinf2 8577.) The main proof is provided by
inf3lema 8559 through inf3lem7 8569, and this final piece eliminates the
auxiliary hypothesis of inf3lem7 8569. This proof is due to
Ian Sutherland, Richard Heck, and Norman Megill and was posted
on Usenet as shown below. Although the result is not new, the authors
were unable to find a published proof.
(As posted to sci.logic on 30-Oct-1996, with annotations added.) Theorem: The statement "There exists a nonempty set that is a subset of its union" implies the Axiom of Infinity. Proof: Let X be a nonempty set which is a subset of its union; the latter property is equivalent to saying that for any y in X, there exists a z in X such that y is in z. Define by finite recursion a function F:omega-->(power X) such that F_0 = 0 (See inf3lemb 8560.) F_n+1 = {y<X | y^X subset F_n} (See inf3lemc 8561.) Note: ^ means intersect, < means \in ("element of"). (Finite recursion as typically done requires the existence of omega; to avoid this we can just use transfinite recursion restricted to omega. F is a class-term that is not necessarily a set at this point.) Lemma 1. F_n subset F_n+1. (See inf3lem1 8563.) Proof: By induction: F_0 subset F_1. If y < F_n+1, then y^X subset F_n, so if F_n subset F_n+1, then y^X subset F_n+1, so y < F_n+2. Lemma 2. F_n =/= X. (See inf3lem2 8564.) Proof: By induction: F_0 =/= X because X is not empty. Assume F_n =/= X. Then there is a y in X that is not in F_n. By definition of X, there is a z in X that contains y. Suppose F_n+1 = X. Then z is in F_n+1, and z^X contains y, so z^X is not a subset of F_n, contrary to the definition of F_n+1. Lemma 3. F_n =/= F_n+1. (See inf3lem3 8565.) Proof: Using the identity y^X subset F_n <-> y^(X-F_n) = 0, we have F_n+1 = {y<X | y^(X-F_n) = 0}. Let q = {y<X-F_n | y^(X-F_n) = 0}. Then q subset F_n+1. Since X-F_n is not empty by Lemma 2 and q is the set of \in-minimal elements of X-F_n, by Foundation q is not empty, so q and therefore F_n+1 have an element not in F_n. Lemma 4. F_n proper_subset F_n+1. (See inf3lem4 8566.) Proof: Lemmas 1 and 3. Lemma 5. F_m proper_subset F_n, m < n. (See inf3lem5 8567.) Proof: Fix m and use induction on n > m. Basis: F_m proper_subset F_m+1 by Lemma 4. Induction: Assume F_m proper_subset F_n. Then since F_n proper_subset F_n+1, F_m proper_subset F_n+1 by transitivity of proper subset. By Lemma 5, F_m =/= F_n for m =/= n, so F is 1-1. (See inf3lem6 8568.) Thus, the inverse of F is a function with range omega and domain a subset of power X, so omega exists by Replacement. (See inf3lem7 8569.) Q.E.D.(Contributed by NM, 29-Oct-1996.) |
⊢ ∃𝑥(𝑥 ≠ ∅ ∧ 𝑥 ⊆ ∪ 𝑥) ⇒ ⊢ ω ∈ V | ||
Theorem | infeq5i 8571 | Half of infeq5 8572. (Contributed by Mario Carneiro, 16-Nov-2014.) |
⊢ (ω ∈ V → ∃𝑥 𝑥 ⊊ ∪ 𝑥) | ||
Theorem | infeq5 8572 | The statement "there exists a set that is a proper subset of its union" is equivalent to the Axiom of Infinity (shown on the right-hand side in the form of omex 8578.) The left-hand side provides us with a very short way to express the Axiom of Infinity using only elementary symbols. This proof of equivalence does not depend on the Axiom of Infinity. (Contributed by NM, 23-Mar-2004.) (Revised by Mario Carneiro, 16-Nov-2014.) |
⊢ (∃𝑥 𝑥 ⊊ ∪ 𝑥 ↔ ω ∈ V) | ||
Axiom | ax-inf 8573* |
Axiom of Infinity. An axiom of Zermelo-Fraenkel set theory. This axiom
is the gateway to "Cantor's paradise" (an expression coined by
Hilbert).
It asserts that given a starting set 𝑥, an infinite set 𝑦 built
from it exists. Although our version is apparently not given in the
literature, it is similar to, but slightly shorter than, the Axiom of
Infinity in [FreydScedrov] p. 283
(see inf1 8557 and inf2 8558). More
standard versions, which essentially state that there exists a set
containing all the natural numbers, are shown as zfinf2 8577 and omex 8578 and
are based on the (nontrivial) proof of inf3 8570.
This version has the
advantage that when expanded to primitives, it has fewer symbols than
the standard version ax-inf2 8576. Theorem inf0 8556
shows the reverse
derivation of our axiom from a standard one. Theorem inf5 8580
shows a
very short way to state this axiom.
The standard version of Infinity ax-inf2 8576 requires this axiom along with Regularity ax-reg 8538 for its derivation (as theorem axinf2 8575 below). In order to more easily identify the normal uses of Regularity, we will usually reference ax-inf2 8576 instead of this one. The derivation of this axiom from ax-inf2 8576 is shown by theorem axinf 8579. Proofs should normally use the standard version ax-inf2 8576 instead of this axiom. (New usage is discouraged.) (Contributed by NM, 16-Aug-1993.) |
⊢ ∃𝑦(𝑥 ∈ 𝑦 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ∃𝑤(𝑧 ∈ 𝑤 ∧ 𝑤 ∈ 𝑦))) | ||
Theorem | zfinf 8574* | Axiom of Infinity expressed with the fewest number of different variables. (New usage is discouraged.) (Contributed by NM, 14-Aug-2003.) |
⊢ ∃𝑥(𝑦 ∈ 𝑥 ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑦 ∈ 𝑧 ∧ 𝑧 ∈ 𝑥))) | ||
Theorem | axinf2 8575* |
A standard version of Axiom of Infinity, expanded to primitives, derived
from our version of Infinity ax-inf 8573 and Regularity ax-reg 8538.
This theorem should not be referenced in any proof. Instead, use ax-inf2 8576 below so that the ordinary uses of Regularity can be more easily identified. (New usage is discouraged.) (Contributed by NM, 3-Nov-1996.) |
⊢ ∃𝑥(∃𝑦(𝑦 ∈ 𝑥 ∧ ∀𝑧 ¬ 𝑧 ∈ 𝑦) ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑧 ∈ 𝑥 ∧ ∀𝑤(𝑤 ∈ 𝑧 ↔ (𝑤 ∈ 𝑦 ∨ 𝑤 = 𝑦))))) | ||
Axiom | ax-inf2 8576* | A standard version of Axiom of Infinity of ZF set theory. In English, it says: there exists a set that contains the empty set and the successors of all of its members. Theorem zfinf2 8577 shows it converted to abbreviations. This axiom was derived as theorem axinf2 8575 above, using our version of Infinity ax-inf 8573 and the Axiom of Regularity ax-reg 8538. We will reference ax-inf2 8576 instead of axinf2 8575 so that the ordinary uses of Regularity can be more easily identified. The reverse derivation of ax-inf 8573 from ax-inf2 8576 is shown by theorem axinf 8579. (Contributed by NM, 3-Nov-1996.) |
⊢ ∃𝑥(∃𝑦(𝑦 ∈ 𝑥 ∧ ∀𝑧 ¬ 𝑧 ∈ 𝑦) ∧ ∀𝑦(𝑦 ∈ 𝑥 → ∃𝑧(𝑧 ∈ 𝑥 ∧ ∀𝑤(𝑤 ∈ 𝑧 ↔ (𝑤 ∈ 𝑦 ∨ 𝑤 = 𝑦))))) | ||
Theorem | zfinf2 8577* | A standard version of the Axiom of Infinity, using definitions to abbreviate. Axiom Inf of [BellMachover] p. 472. (See ax-inf2 8576 for the unabbreviated version.) (Contributed by NM, 30-Aug-1993.) |
⊢ ∃𝑥(∅ ∈ 𝑥 ∧ ∀𝑦 ∈ 𝑥 suc 𝑦 ∈ 𝑥) | ||
Theorem | omex 8578 |
The existence of omega (the class of natural numbers). Axiom 7 of
[TakeutiZaring] p. 43. This
theorem is proved assuming the Axiom of
Infinity and in fact is equivalent to it, as shown by the reverse
derivation inf0 8556.
A finitist (someone who doesn't believe in infinity) could, without contradiction, replace the Axiom of Infinity by its denial ¬ ω ∈ V; this would lead to ω = On by omon 7118 and Fin = V (the universe of all sets) by fineqv 8216. The finitist could still develop natural number, integer, and rational number arithmetic but would be denied the real numbers (as well as much of the rest of mathematics). In deference to the finitist, much of our development is done, when possible, without invoking the Axiom of Infinity; an example is Peano's axioms peano1 7127 through peano5 7131 (which many textbooks prove more easily assuming Infinity). (Contributed by NM, 6-Aug-1994.) |
⊢ ω ∈ V | ||
Theorem | axinf 8579* | The first version of the Axiom of Infinity ax-inf 8573 proved from the second version ax-inf2 8576. Note that we didn't use ax-reg 8538, unlike the other direction axinf2 8575. (Contributed by NM, 24-Apr-2009.) |
⊢ ∃𝑦(𝑥 ∈ 𝑦 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ∃𝑤(𝑧 ∈ 𝑤 ∧ 𝑤 ∈ 𝑦))) | ||
Theorem | inf5 8580 | The statement "there exists a set that is a proper subset of its union" is equivalent to the Axiom of Infinity (see theorem infeq5 8572). This provides us with a very compact way to express the Axiom of Infinity using only elementary symbols. (Contributed by NM, 3-Jun-2005.) |
⊢ ∃𝑥 𝑥 ⊊ ∪ 𝑥 | ||
Theorem | omelon 8581 | Omega is an ordinal number. (Contributed by NM, 10-May-1998.) (Revised by Mario Carneiro, 30-Jan-2013.) |
⊢ ω ∈ On | ||
Theorem | dfom3 8582* | The class of natural numbers omega can be defined as the smallest "inductive set," which is valid provided we assume the Axiom of Infinity. Definition 6.3 of [Eisenberg] p. 82. (Contributed by NM, 6-Aug-1994.) |
⊢ ω = ∩ {𝑥 ∣ (∅ ∈ 𝑥 ∧ ∀𝑦 ∈ 𝑥 suc 𝑦 ∈ 𝑥)} | ||
Theorem | elom3 8583* | A simplification of elom 7110 assuming the Axiom of Infinity. (Contributed by NM, 30-May-2003.) |
⊢ (𝐴 ∈ ω ↔ ∀𝑥(Lim 𝑥 → 𝐴 ∈ 𝑥)) | ||
Theorem | dfom4 8584* | A simplification of df-om 7108 assuming the Axiom of Infinity. (Contributed by NM, 30-May-2003.) |
⊢ ω = {𝑥 ∣ ∀𝑦(Lim 𝑦 → 𝑥 ∈ 𝑦)} | ||
Theorem | dfom5 8585 | ω is the smallest limit ordinal and can be defined as such (although the Axiom of Infinity is needed to ensure that at least one limit ordinal exists). (Contributed by FL, 22-Feb-2011.) (Revised by Mario Carneiro, 2-Feb-2013.) |
⊢ ω = ∩ {𝑥 ∣ Lim 𝑥} | ||
Theorem | oancom 8586 | Ordinal addition is not commutative. This theorem shows a counterexample. Remark in [TakeutiZaring] p. 60. (Contributed by NM, 10-Dec-2004.) |
⊢ (1𝑜 +𝑜 ω) ≠ (ω +𝑜 1𝑜) | ||
Theorem | isfinite 8587 | A set is finite iff it is strictly dominated by the class of natural number. Theorem 42 of [Suppes] p. 151. The Axiom of Infinity is used for the forward implication. (Contributed by FL, 16-Apr-2011.) |
⊢ (𝐴 ∈ Fin ↔ 𝐴 ≺ ω) | ||
Theorem | fict 8588 | A finite set is countable (weaker version of isfinite 8587). (Contributed by Thierry Arnoux, 27-Mar-2018.) |
⊢ (𝐴 ∈ Fin → 𝐴 ≼ ω) | ||
Theorem | nnsdom 8589 | A natural number is strictly dominated by the set of natural numbers. Example 3 of [Enderton] p. 146. (Contributed by NM, 28-Oct-2003.) |
⊢ (𝐴 ∈ ω → 𝐴 ≺ ω) | ||
Theorem | omenps 8590 | Omega is equinumerous to a proper subset of itself. Example 13.2(4) of [Eisenberg] p. 216. (Contributed by NM, 30-Jul-2003.) |
⊢ ω ≈ (ω ∖ {∅}) | ||
Theorem | omensuc 8591 | The set of natural numbers is equinumerous to its successor. (Contributed by NM, 30-Oct-2003.) |
⊢ ω ≈ suc ω | ||
Theorem | infdifsn 8592 | Removing a singleton from an infinite set does not change the cardinality of the set. (Contributed by Mario Carneiro, 30-Apr-2015.) (Revised by Mario Carneiro, 16-May-2015.) |
⊢ (ω ≼ 𝐴 → (𝐴 ∖ {𝐵}) ≈ 𝐴) | ||
Theorem | infdiffi 8593 | Removing a finite set from an infinite set does not change the cardinality of the set. (Contributed by Mario Carneiro, 30-Apr-2015.) |
⊢ ((ω ≼ 𝐴 ∧ 𝐵 ∈ Fin) → (𝐴 ∖ 𝐵) ≈ 𝐴) | ||
Theorem | unbnn3 8594* | Any unbounded subset of natural numbers is equinumerous to the set of all natural numbers. This version of unbnn 8257 eliminates its hypothesis by assuming the Axiom of Infinity. (Contributed by NM, 4-May-2005.) |
⊢ ((𝐴 ⊆ ω ∧ ∀𝑥 ∈ ω ∃𝑦 ∈ 𝐴 𝑥 ∈ 𝑦) → 𝐴 ≈ ω) | ||
Theorem | noinfep 8595* | Using the Axiom of Regularity in the form zfregfr 8547, show that there are no infinite descending ∈-chains. Proposition 7.34 of [TakeutiZaring] p. 44. (Contributed by NM, 26-Jan-2006.) (Revised by Mario Carneiro, 22-Mar-2013.) |
⊢ ∃𝑥 ∈ ω (𝐹‘suc 𝑥) ∉ (𝐹‘𝑥) | ||
Syntax | ccnf 8596 | Extend class notation with the Cantor normal form function. |
class CNF | ||
Definition | df-cnf 8597* | Define the Cantor normal form function, which takes as input a finitely supported function from 𝑦 to 𝑥 and outputs the corresponding member of the ordinal exponential 𝑥 ↑𝑜 𝑦. The content of the original Cantor Normal Form theorem is that for 𝑥 = ω this function is a bijection onto ω ↑𝑜 𝑦 for any ordinal 𝑦 (or, since the function restricts naturally to different ordinals, the statement that the composite function is a bijection to On). More can be said about the function, however, and in particular it is an order isomorphism for a certain easily defined well-ordering of the finitely supported functions, which gives an alternate definition cantnffval2 8630 of this function in terms of df-oi 8456. (Contributed by Mario Carneiro, 25-May-2015.) (Revised by AV, 28-Jun-2019.) |
⊢ CNF = (𝑥 ∈ On, 𝑦 ∈ On ↦ (𝑓 ∈ {𝑔 ∈ (𝑥 ↑𝑚 𝑦) ∣ 𝑔 finSupp ∅} ↦ ⦋OrdIso( E , (𝑓 supp ∅)) / ℎ⦌(seq𝜔((𝑘 ∈ V, 𝑧 ∈ V ↦ (((𝑥 ↑𝑜 (ℎ‘𝑘)) ·𝑜 (𝑓‘(ℎ‘𝑘))) +𝑜 𝑧)), ∅)‘dom ℎ))) | ||
Theorem | cantnffval 8598* | The value of the Cantor normal form function. (Contributed by Mario Carneiro, 25-May-2015.) (Revised by AV, 28-Jun-2019.) |
⊢ 𝑆 = {𝑔 ∈ (𝐴 ↑𝑚 𝐵) ∣ 𝑔 finSupp ∅} & ⊢ (𝜑 → 𝐴 ∈ On) & ⊢ (𝜑 → 𝐵 ∈ On) ⇒ ⊢ (𝜑 → (𝐴 CNF 𝐵) = (𝑓 ∈ 𝑆 ↦ ⦋OrdIso( E , (𝑓 supp ∅)) / ℎ⦌(seq𝜔((𝑘 ∈ V, 𝑧 ∈ V ↦ (((𝐴 ↑𝑜 (ℎ‘𝑘)) ·𝑜 (𝑓‘(ℎ‘𝑘))) +𝑜 𝑧)), ∅)‘dom ℎ))) | ||
Theorem | cantnfdm 8599* | The domain of the Cantor normal form function (in later lemmas we will use dom (𝐴 CNF 𝐵) to abbreviate "the set of finitely supported functions from 𝐵 to 𝐴"). (Contributed by Mario Carneiro, 25-May-2015.) (Revised by AV, 28-Jun-2019.) |
⊢ 𝑆 = {𝑔 ∈ (𝐴 ↑𝑚 𝐵) ∣ 𝑔 finSupp ∅} & ⊢ (𝜑 → 𝐴 ∈ On) & ⊢ (𝜑 → 𝐵 ∈ On) ⇒ ⊢ (𝜑 → dom (𝐴 CNF 𝐵) = 𝑆) | ||
Theorem | cantnfvalf 8600* | Lemma for cantnf 8628. The function appearing in cantnfval 8603 is unconditionally a function. (Contributed by Mario Carneiro, 20-May-2015.) |
⊢ 𝐹 = seq𝜔((𝑘 ∈ 𝐴, 𝑧 ∈ 𝐵 ↦ (𝐶 +𝑜 𝐷)), ∅) ⇒ ⊢ 𝐹:ω⟶On |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |