Home | Intuitionistic Logic Explorer Theorem List (p. 46 of 144) | < 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 | onsucelsucr 4501 | Membership is inherited by predecessors. The converse, for all ordinals, implies excluded middle, as shown at onsucelsucexmid 4523. However, the converse does hold where 𝐵 is a natural number, as seen at nnsucelsuc 6482. (Contributed by Jim Kingdon, 17-Jul-2019.) |
⊢ (𝐵 ∈ On → (suc 𝐴 ∈ suc 𝐵 → 𝐴 ∈ 𝐵)) | ||
Theorem | onsucsssucr 4502 | The subclass relationship between two ordinals is inherited by their predecessors. The converse implies excluded middle, as shown at onsucsssucexmid 4520. (Contributed by Mario Carneiro and Jim Kingdon, 29-Jul-2019.) |
⊢ ((𝐴 ∈ On ∧ Ord 𝐵) → (suc 𝐴 ⊆ suc 𝐵 → 𝐴 ⊆ 𝐵)) | ||
Theorem | sucunielr 4503 | Successor and union. The converse (where 𝐵 is an ordinal) implies excluded middle, as seen at ordsucunielexmid 4524. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ (suc 𝐴 ∈ 𝐵 → 𝐴 ∈ ∪ 𝐵) | ||
Theorem | unon 4504 | The class of all ordinal numbers is its own union. Exercise 11 of [TakeutiZaring] p. 40. (Contributed by NM, 12-Nov-2003.) |
⊢ ∪ On = On | ||
Theorem | onuniss2 4505* | The union of the ordinal subsets of an ordinal number is that number. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ (𝐴 ∈ On → ∪ {𝑥 ∈ On ∣ 𝑥 ⊆ 𝐴} = 𝐴) | ||
Theorem | limon 4506 | The class of ordinal numbers is a limit ordinal. (Contributed by NM, 24-Mar-1995.) |
⊢ Lim On | ||
Theorem | ordunisuc2r 4507* | An ordinal which contains the successor of each of its members is equal to its union. (Contributed by Jim Kingdon, 14-Nov-2018.) |
⊢ (Ord 𝐴 → (∀𝑥 ∈ 𝐴 suc 𝑥 ∈ 𝐴 → 𝐴 = ∪ 𝐴)) | ||
Theorem | onssi 4508 | An ordinal number is a subset of On. (Contributed by NM, 11-Aug-1994.) |
⊢ 𝐴 ∈ On ⇒ ⊢ 𝐴 ⊆ On | ||
Theorem | onsuci 4509 | The successor of an ordinal number is an ordinal number. Corollary 7N(c) of [Enderton] p. 193. (Contributed by NM, 12-Jun-1994.) |
⊢ 𝐴 ∈ On ⇒ ⊢ suc 𝐴 ∈ On | ||
Theorem | onintonm 4510* | The intersection of an inhabited collection of ordinal numbers is an ordinal number. Compare Exercise 6 of [TakeutiZaring] p. 44. (Contributed by Mario Carneiro and Jim Kingdon, 30-Aug-2021.) |
⊢ ((𝐴 ⊆ On ∧ ∃𝑥 𝑥 ∈ 𝐴) → ∩ 𝐴 ∈ On) | ||
Theorem | onintrab2im 4511 | An existence condition which implies an intersection is an ordinal number. (Contributed by Jim Kingdon, 30-Aug-2021.) |
⊢ (∃𝑥 ∈ On 𝜑 → ∩ {𝑥 ∈ On ∣ 𝜑} ∈ On) | ||
Theorem | ordtriexmidlem 4512 | Lemma for decidability and ordinals. The set {𝑥 ∈ {∅} ∣ 𝜑} is a way of connecting statements about ordinals (such as trichotomy in ordtriexmid 4514 or weak linearity in ordsoexmid 4555) with a proposition 𝜑. Our lemma states that it is an ordinal number. (Contributed by Jim Kingdon, 28-Jan-2019.) |
⊢ {𝑥 ∈ {∅} ∣ 𝜑} ∈ On | ||
Theorem | ordtriexmidlem2 4513* | Lemma for decidability and ordinals. The set {𝑥 ∈ {∅} ∣ 𝜑} is a way of connecting statements about ordinals (such as trichotomy in ordtriexmid 4514 or weak linearity in ordsoexmid 4555) with a proposition 𝜑. Our lemma helps connect that set to excluded middle. (Contributed by Jim Kingdon, 28-Jan-2019.) |
⊢ ({𝑥 ∈ {∅} ∣ 𝜑} = ∅ → ¬ 𝜑) | ||
Theorem | ordtriexmid 4514* |
Ordinal trichotomy implies the law of the excluded middle (that is,
decidability of an arbitrary proposition).
This theorem is stated in "Constructive ordinals", [Crosilla], p. "Set-theoretic principles incompatible with intuitionistic logic". Also see exmidontri 7228 which is much the same theorem but biconditionalized and using the EXMID notation. (Contributed by Mario Carneiro and Jim Kingdon, 14-Nov-2018.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ontriexmidim 4515* | Ordinal trichotomy implies excluded middle. Closed form of ordtriexmid 4514. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥) → DECID 𝜑) | ||
Theorem | ordtri2orexmid 4516* | Ordinal trichotomy implies excluded middle. (Contributed by Jim Kingdon, 31-Jul-2019.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑦 ⊆ 𝑥) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | 2ordpr 4517 | Version of 2on 6416 with the definition of 2o expanded and expressed in terms of Ord. (Contributed by Jim Kingdon, 29-Aug-2021.) |
⊢ Ord {∅, {∅}} | ||
Theorem | ontr2exmid 4518* | An ordinal transitivity law which implies excluded middle. (Contributed by Jim Kingdon, 17-Sep-2021.) |
⊢ ∀𝑥 ∈ On ∀𝑦∀𝑧 ∈ On ((𝑥 ⊆ 𝑦 ∧ 𝑦 ∈ 𝑧) → 𝑥 ∈ 𝑧) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ordtri2or2exmidlem 4519* | A set which is 2o if 𝜑 or ∅ if ¬ 𝜑 is an ordinal. (Contributed by Jim Kingdon, 29-Aug-2021.) |
⊢ {𝑥 ∈ {∅, {∅}} ∣ 𝜑} ∈ On | ||
Theorem | onsucsssucexmid 4520* | The converse of onsucsssucr 4502 implies excluded middle. (Contributed by Mario Carneiro and Jim Kingdon, 29-Jul-2019.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 → suc 𝑥 ⊆ suc 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | onsucelsucexmidlem1 4521* | Lemma for onsucelsucexmid 4523. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ ∅ ∈ {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = ∅ ∨ 𝜑)} | ||
Theorem | onsucelsucexmidlem 4522* | Lemma for onsucelsucexmid 4523. The set {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = ∅ ∨ 𝜑)} appears as 𝐴 in the proof of Theorem 1.3 in [Bauer] p. 483 (see acexmidlema 5856), and similar sets also appear in other proofs that various propositions imply excluded middle, for example in ordtriexmidlem 4512. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = ∅ ∨ 𝜑)} ∈ On | ||
Theorem | onsucelsucexmid 4523* | The converse of onsucelsucr 4501 implies excluded middle. On the other hand, if 𝑦 is constrained to be a natural number, instead of an arbitrary ordinal, then the converse of onsucelsucr 4501 does hold, as seen at nnsucelsuc 6482. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 → suc 𝑥 ∈ suc 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ordsucunielexmid 4524* | The converse of sucunielr 4503 (where 𝐵 is an ordinal) implies excluded middle. (Contributed by Jim Kingdon, 2-Aug-2019.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ ∪ 𝑦 → suc 𝑥 ∈ 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | regexmidlemm 4525* | Lemma for regexmid 4528. 𝐴 is inhabited. (Contributed by Jim Kingdon, 3-Sep-2019.) |
⊢ 𝐴 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = {∅} ∨ (𝑥 = ∅ ∧ 𝜑))} ⇒ ⊢ ∃𝑦 𝑦 ∈ 𝐴 | ||
Theorem | regexmidlem1 4526* | Lemma for regexmid 4528. If 𝐴 has a minimal element, excluded middle follows. (Contributed by Jim Kingdon, 3-Sep-2019.) |
⊢ 𝐴 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = {∅} ∨ (𝑥 = ∅ ∧ 𝜑))} ⇒ ⊢ (∃𝑦(𝑦 ∈ 𝐴 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ¬ 𝑧 ∈ 𝐴)) → (𝜑 ∨ ¬ 𝜑)) | ||
Theorem | reg2exmidlema 4527* | Lemma for reg2exmid 4529. If 𝐴 has a minimal element (expressed by ⊆), excluded middle follows. (Contributed by Jim Kingdon, 2-Oct-2021.) |
⊢ 𝐴 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = {∅} ∨ (𝑥 = ∅ ∧ 𝜑))} ⇒ ⊢ (∃𝑢 ∈ 𝐴 ∀𝑣 ∈ 𝐴 𝑢 ⊆ 𝑣 → (𝜑 ∨ ¬ 𝜑)) | ||
Theorem | regexmid 4528* |
The axiom of foundation implies excluded middle.
By foundation (or regularity), we mean the principle that every inhabited set has an element which is minimal (when arranged by ∈). The statement of foundation here is taken from Metamath Proof Explorer's ax-reg, and is identical (modulo one unnecessary quantifier) to the statement of foundation in Theorem "Foundation implies instances of EM" of [Crosilla], p. "Set-theoretic principles incompatible with intuitionistic logic". For this reason, IZF does not adopt foundation as an axiom and instead replaces it with ax-setind 4530. (Contributed by Jim Kingdon, 3-Sep-2019.) |
⊢ (∃𝑦 𝑦 ∈ 𝑥 → ∃𝑦(𝑦 ∈ 𝑥 ∧ ∀𝑧(𝑧 ∈ 𝑦 → ¬ 𝑧 ∈ 𝑥))) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | reg2exmid 4529* | If any inhabited set has a minimal element (when expressed by ⊆), excluded middle follows. (Contributed by Jim Kingdon, 2-Oct-2021.) |
⊢ ∀𝑧(∃𝑤 𝑤 ∈ 𝑧 → ∃𝑥 ∈ 𝑧 ∀𝑦 ∈ 𝑧 𝑥 ⊆ 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Axiom | ax-setind 4530* |
Axiom of ∈-Induction (also known as set
induction). An axiom of
Intuitionistic Zermelo-Fraenkel set theory. Axiom 9 of [Crosilla] p.
"Axioms of CZF and IZF". This replaces the Axiom of
Foundation (also
called Regularity) from Zermelo-Fraenkel set theory.
For more on axioms which might be adopted which are incompatible with this axiom (that is, Non-wellfounded Set Theory but in the absence of excluded middle), see Chapter 20 of [AczelRathjen], p. 183. (Contributed by Jim Kingdon, 19-Oct-2018.) |
⊢ (∀𝑎(∀𝑦 ∈ 𝑎 [𝑦 / 𝑎]𝜑 → 𝜑) → ∀𝑎𝜑) | ||
Theorem | setindel 4531* | ∈-Induction in terms of membership in a class. (Contributed by Mario Carneiro and Jim Kingdon, 22-Oct-2018.) |
⊢ (∀𝑥(∀𝑦(𝑦 ∈ 𝑥 → 𝑦 ∈ 𝑆) → 𝑥 ∈ 𝑆) → 𝑆 = V) | ||
Theorem | setind 4532* | Set (epsilon) induction. Theorem 5.22 of [TakeutiZaring] p. 21. (Contributed by NM, 17-Sep-2003.) |
⊢ (∀𝑥(𝑥 ⊆ 𝐴 → 𝑥 ∈ 𝐴) → 𝐴 = V) | ||
Theorem | setind2 4533 | Set (epsilon) induction, stated compactly. Given as a homework problem in 1992 by George Boolos (1940-1996). (Contributed by NM, 17-Sep-2003.) |
⊢ (𝒫 𝐴 ⊆ 𝐴 → 𝐴 = V) | ||
Theorem | elirr 4534 |
No class is a member of itself. Exercise 6 of [TakeutiZaring] p. 22.
The reason that this theorem is marked as discouraged is a bit subtle. If we wanted to reduce usage of ax-setind 4530, we could redefine Ord 𝐴 (df-iord 4360) to also require E Fr 𝐴 (df-frind 4326) and in that case any theorem related to irreflexivity of ordinals could use ordirr 4535 (which under that definition would presumably not need ax-setind 4530 to prove it). But since ordinals have not yet been defined that way, we cannot rely on the "don't add additional axiom use" feature of the minimizer to get theorems to use ordirr 4535. To encourage ordirr 4535 when possible, we mark this theorem as discouraged. (Contributed by NM, 7-Aug-1994.) (Proof rewritten by Mario Carneiro and Jim Kingdon, 26-Nov-2018.) (New usage is discouraged.) |
⊢ ¬ 𝐴 ∈ 𝐴 | ||
Theorem | ordirr 4535 | Epsilon irreflexivity of ordinals: no ordinal class is a member of itself. Theorem 2.2(i) of [BellMachover] p. 469, generalized to classes. The present proof requires ax-setind 4530. If in the definition of ordinals df-iord 4360, we also required that membership be well-founded on any ordinal (see df-frind 4326), then we could prove ordirr 4535 without ax-setind 4530. (Contributed by NM, 2-Jan-1994.) |
⊢ (Ord 𝐴 → ¬ 𝐴 ∈ 𝐴) | ||
Theorem | onirri 4536 | An ordinal number is not a member of itself. Theorem 7M(c) of [Enderton] p. 192. (Contributed by NM, 11-Jun-1994.) |
⊢ 𝐴 ∈ On ⇒ ⊢ ¬ 𝐴 ∈ 𝐴 | ||
Theorem | nordeq 4537 | A member of an ordinal class is not equal to it. (Contributed by NM, 25-May-1998.) |
⊢ ((Ord 𝐴 ∧ 𝐵 ∈ 𝐴) → 𝐴 ≠ 𝐵) | ||
Theorem | ordn2lp 4538 | An ordinal class cannot be an element of one of its members. Variant of first part of Theorem 2.2(vii) of [BellMachover] p. 469. (Contributed by NM, 3-Apr-1994.) |
⊢ (Ord 𝐴 → ¬ (𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐴)) | ||
Theorem | orddisj 4539 | An ordinal class and its singleton are disjoint. (Contributed by NM, 19-May-1998.) |
⊢ (Ord 𝐴 → (𝐴 ∩ {𝐴}) = ∅) | ||
Theorem | orddif 4540 | Ordinal derived from its successor. (Contributed by NM, 20-May-1998.) |
⊢ (Ord 𝐴 → 𝐴 = (suc 𝐴 ∖ {𝐴})) | ||
Theorem | elirrv 4541 | The membership relation is irreflexive: no set is a member of itself. Theorem 105 of [Suppes] p. 54. (Contributed by NM, 19-Aug-1993.) |
⊢ ¬ 𝑥 ∈ 𝑥 | ||
Theorem | sucprcreg 4542 | A class is equal to its successor iff it is a proper class (assuming the Axiom of Set Induction). (Contributed by NM, 9-Jul-2004.) |
⊢ (¬ 𝐴 ∈ V ↔ suc 𝐴 = 𝐴) | ||
Theorem | ruv 4543 | 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 4544 | Alternate proof of Russell's Paradox ru 2959, simplified using (indirectly) the Axiom of Set Induction ax-setind 4530. (Contributed by Alan Sare, 4-Oct-2008.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ {𝑥 ∣ 𝑥 ∉ 𝑥} ∉ V | ||
Theorem | onprc 4545 | No set contains all ordinal numbers. Proposition 7.13 of [TakeutiZaring] p. 38. This is also known as the Burali-Forti paradox (remark in [Enderton] p. 194). In 1897, Cesare Burali-Forti noticed that since the "set" of all ordinal numbers is an ordinal class (ordon 4479), it must be both an element of the set of all ordinal numbers yet greater than every such element. ZF set theory resolves this paradox by not allowing the class of all ordinal numbers to be a set (so instead it is a proper class). Here we prove the denial of its existence. (Contributed by NM, 18-May-1994.) |
⊢ ¬ On ∈ V | ||
Theorem | sucon 4546 | The class of all ordinal numbers is its own successor. (Contributed by NM, 12-Sep-2003.) |
⊢ suc On = On | ||
Theorem | en2lp 4547 | No class has 2-cycle membership loops. Theorem 7X(b) of [Enderton] p. 206. (Contributed by NM, 16-Oct-1996.) (Proof rewritten by Mario Carneiro and Jim Kingdon, 27-Nov-2018.) |
⊢ ¬ (𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝐴) | ||
Theorem | preleq 4548 | 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 4549 | Theorem for alternate representation of ordered pairs, requiring the Axiom of Set Induction ax-setind 4530 (via the preleq 4548 step). See df-op 3598 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 | suc11g 4550 | The successor operation behaves like a one-to-one function (assuming the Axiom of Set Induction). Similar to Exercise 35 of [Enderton] p. 208 and its converse. (Contributed by NM, 25-Oct-2003.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → (suc 𝐴 = suc 𝐵 ↔ 𝐴 = 𝐵)) | ||
Theorem | suc11 4551 | The successor operation behaves like a one-to-one function. Compare Exercise 16 of [Enderton] p. 194. (Contributed by NM, 3-Sep-2003.) |
⊢ ((𝐴 ∈ On ∧ 𝐵 ∈ On) → (suc 𝐴 = suc 𝐵 ↔ 𝐴 = 𝐵)) | ||
Theorem | dtruex 4552* | At least two sets exist (or in terms of first-order logic, the universe of discourse has two or more objects). Although dtruarb 4186 can also be summarized as "at least two sets exist", the difference is that dtruarb 4186 shows the existence of two sets which are not equal to each other, but this theorem says that given a specific 𝑦, we can construct a set 𝑥 which does not equal it. (Contributed by Jim Kingdon, 29-Dec-2018.) |
⊢ ∃𝑥 ¬ 𝑥 = 𝑦 | ||
Theorem | dtru 4553* | At least two sets exist (or in terms of first-order logic, the universe of discourse has two or more objects). If we assumed the law of the excluded middle this would be equivalent to dtruex 4552. (Contributed by Jim Kingdon, 29-Dec-2018.) |
⊢ ¬ ∀𝑥 𝑥 = 𝑦 | ||
Theorem | eunex 4554 | Existential uniqueness implies there is a value for which the wff argument is false. (Contributed by Jim Kingdon, 29-Dec-2018.) |
⊢ (∃!𝑥𝜑 → ∃𝑥 ¬ 𝜑) | ||
Theorem | ordsoexmid 4555 | Weak linearity of ordinals implies the law of the excluded middle (that is, decidability of an arbitrary proposition). (Contributed by Mario Carneiro and Jim Kingdon, 29-Jan-2019.) |
⊢ E Or On ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ordsuc 4556 | The successor of an ordinal class is ordinal. (Contributed by NM, 3-Apr-1995.) (Constructive proof by Mario Carneiro and Jim Kingdon, 20-Jul-2019.) |
⊢ (Ord 𝐴 ↔ Ord suc 𝐴) | ||
Theorem | onsucuni2 4557 | A successor ordinal is the successor of its union. (Contributed by NM, 10-Dec-2004.) (Proof shortened by Andrew Salmon, 27-Aug-2011.) |
⊢ ((𝐴 ∈ On ∧ 𝐴 = suc 𝐵) → suc ∪ 𝐴 = 𝐴) | ||
Theorem | 0elsucexmid 4558* | If the successor of any ordinal class contains the empty set, excluded middle follows. (Contributed by Jim Kingdon, 3-Sep-2021.) |
⊢ ∀𝑥 ∈ On ∅ ∈ suc 𝑥 ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | nlimsucg 4559 | A successor is not a limit ordinal. (Contributed by NM, 25-Mar-1995.) (Proof shortened by Andrew Salmon, 27-Aug-2011.) |
⊢ (𝐴 ∈ 𝑉 → ¬ Lim suc 𝐴) | ||
Theorem | ordpwsucss 4560 |
The collection of ordinals in the power class of an ordinal is a
superset of its successor.
We can think of (𝒫 𝐴 ∩ On) as another possible definition of successor, which would be equivalent to df-suc 4365 given excluded middle. It is an ordinal, and has some successor-like properties. For example, if 𝐴 ∈ On then both ∪ suc 𝐴 = 𝐴 (onunisuci 4426) and ∪ {𝑥 ∈ On ∣ 𝑥 ⊆ 𝐴} = 𝐴 (onuniss2 4505). Constructively (𝒫 𝐴 ∩ On) and suc 𝐴 cannot be shown to be equivalent (as proved at ordpwsucexmid 4563). (Contributed by Jim Kingdon, 21-Jul-2019.) |
⊢ (Ord 𝐴 → suc 𝐴 ⊆ (𝒫 𝐴 ∩ On)) | ||
Theorem | onnmin 4561 | No member of a set of ordinal numbers belongs to its minimum. (Contributed by NM, 2-Feb-1997.) (Constructive proof by Mario Carneiro and Jim Kingdon, 21-Jul-2019.) |
⊢ ((𝐴 ⊆ On ∧ 𝐵 ∈ 𝐴) → ¬ 𝐵 ∈ ∩ 𝐴) | ||
Theorem | ssnel 4562 | Relationship between subset and elementhood. In the context of ordinals this can be seen as an ordering law. (Contributed by Jim Kingdon, 22-Jul-2019.) |
⊢ (𝐴 ⊆ 𝐵 → ¬ 𝐵 ∈ 𝐴) | ||
Theorem | ordpwsucexmid 4563* | The subset in ordpwsucss 4560 cannot be equality. That is, strengthening it to equality implies excluded middle. (Contributed by Jim Kingdon, 30-Jul-2019.) |
⊢ ∀𝑥 ∈ On suc 𝑥 = (𝒫 𝑥 ∩ On) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ordtri2or2exmid 4564* | Ordinal trichotomy implies excluded middle. (Contributed by Jim Kingdon, 29-Aug-2021.) |
⊢ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | ontri2orexmidim 4565* | Ordinal trichotomy implies excluded middle. Closed form of ordtri2or2exmid 4564. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥) → DECID 𝜑) | ||
Theorem | onintexmid 4566* | If the intersection (infimum) of an inhabited class of ordinal numbers belongs to the class, excluded middle follows. The hypothesis would be provable given excluded middle. (Contributed by Mario Carneiro and Jim Kingdon, 29-Aug-2021.) |
⊢ ((𝑦 ⊆ On ∧ ∃𝑥 𝑥 ∈ 𝑦) → ∩ 𝑦 ∈ 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | zfregfr 4567 | The epsilon relation is well-founded on any class. (Contributed by NM, 26-Nov-1995.) |
⊢ E Fr 𝐴 | ||
Theorem | ordfr 4568 | Epsilon is well-founded on an ordinal class. (Contributed by NM, 22-Apr-1994.) |
⊢ (Ord 𝐴 → E Fr 𝐴) | ||
Theorem | ordwe 4569 | Epsilon well-orders every ordinal. Proposition 7.4 of [TakeutiZaring] p. 36. (Contributed by NM, 3-Apr-1994.) |
⊢ (Ord 𝐴 → E We 𝐴) | ||
Theorem | wetriext 4570* | A trichotomous well-order is extensional. (Contributed by Jim Kingdon, 26-Sep-2021.) |
⊢ (𝜑 → 𝑅 We 𝐴) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → ∀𝑎 ∈ 𝐴 ∀𝑏 ∈ 𝐴 (𝑎𝑅𝑏 ∨ 𝑎 = 𝑏 ∨ 𝑏𝑅𝑎)) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑧 ∈ 𝐴 (𝑧𝑅𝐵 ↔ 𝑧𝑅𝐶)) ⇒ ⊢ (𝜑 → 𝐵 = 𝐶) | ||
Theorem | wessep 4571 | A subset of a set well-ordered by set membership is well-ordered by set membership. (Contributed by Jim Kingdon, 30-Sep-2021.) |
⊢ (( E We 𝐴 ∧ 𝐵 ⊆ 𝐴) → E We 𝐵) | ||
Theorem | reg3exmidlemwe 4572* | Lemma for reg3exmid 4573. Our counterexample 𝐴 satisfies We. (Contributed by Jim Kingdon, 3-Oct-2021.) |
⊢ 𝐴 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = {∅} ∨ (𝑥 = ∅ ∧ 𝜑))} ⇒ ⊢ E We 𝐴 | ||
Theorem | reg3exmid 4573* | If any inhabited set satisfying df-wetr 4328 for E has a minimal element, excluded middle follows. (Contributed by Jim Kingdon, 3-Oct-2021.) |
⊢ (( E We 𝑧 ∧ ∃𝑤 𝑤 ∈ 𝑧) → ∃𝑥 ∈ 𝑧 ∀𝑦 ∈ 𝑧 𝑥 ⊆ 𝑦) ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
Theorem | dcextest 4574* | If it is decidable whether {𝑥 ∣ 𝜑} is a set, then ¬ 𝜑 is decidable (where 𝑥 does not occur in 𝜑). From this fact, we can deduce (outside the formal system, since we cannot quantify over classes) that if it is decidable whether any class is a set, then "weak excluded middle" (that is, any negated proposition ¬ 𝜑 is decidable) holds. (Contributed by Jim Kingdon, 3-Jul-2022.) |
⊢ DECID {𝑥 ∣ 𝜑} ∈ V ⇒ ⊢ DECID ¬ 𝜑 | ||
Theorem | tfi 4575* |
The Principle of Transfinite Induction. Theorem 7.17 of [TakeutiZaring]
p. 39. This principle states that if 𝐴 is a class of ordinal
numbers with the property that every ordinal number included in 𝐴
also belongs to 𝐴, then every ordinal number is in
𝐴.
(Contributed by NM, 18-Feb-2004.) |
⊢ ((𝐴 ⊆ On ∧ ∀𝑥 ∈ On (𝑥 ⊆ 𝐴 → 𝑥 ∈ 𝐴)) → 𝐴 = On) | ||
Theorem | tfis 4576* | Transfinite Induction Schema. If all ordinal numbers less than a given number 𝑥 have a property (induction hypothesis), then all ordinal numbers have the property (conclusion). Exercise 25 of [Enderton] p. 200. (Contributed by NM, 1-Aug-1994.) (Revised by Mario Carneiro, 20-Nov-2016.) |
⊢ (𝑥 ∈ On → (∀𝑦 ∈ 𝑥 [𝑦 / 𝑥]𝜑 → 𝜑)) ⇒ ⊢ (𝑥 ∈ On → 𝜑) | ||
Theorem | tfis2f 4577* | Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 18-Aug-1994.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 ∈ On → (∀𝑦 ∈ 𝑥 𝜓 → 𝜑)) ⇒ ⊢ (𝑥 ∈ On → 𝜑) | ||
Theorem | tfis2 4578* | Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 18-Aug-1994.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 ∈ On → (∀𝑦 ∈ 𝑥 𝜓 → 𝜑)) ⇒ ⊢ (𝑥 ∈ On → 𝜑) | ||
Theorem | tfis3 4579* | Transfinite Induction Schema, using implicit substitution. (Contributed by NM, 4-Nov-2003.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 ∈ On → (∀𝑦 ∈ 𝑥 𝜓 → 𝜑)) ⇒ ⊢ (𝐴 ∈ On → 𝜒) | ||
Theorem | tfisi 4580* | A transfinite induction scheme in "implicit" form where the induction is done on an object derived from the object of interest. (Contributed by Stefan O'Rear, 24-Aug-2015.) |
⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝑇 ∈ On) & ⊢ ((𝜑 ∧ (𝑅 ∈ On ∧ 𝑅 ⊆ 𝑇) ∧ ∀𝑦(𝑆 ∈ 𝑅 → 𝜒)) → 𝜓) & ⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) & ⊢ (𝑥 = 𝐴 → (𝜓 ↔ 𝜃)) & ⊢ (𝑥 = 𝑦 → 𝑅 = 𝑆) & ⊢ (𝑥 = 𝐴 → 𝑅 = 𝑇) ⇒ ⊢ (𝜑 → 𝜃) | ||
Axiom | ax-iinf 4581* | Axiom of Infinity. Axiom 5 of [Crosilla] p. "Axioms of CZF and IZF". (Contributed by Jim Kingdon, 16-Nov-2018.) |
⊢ ∃𝑥(∅ ∈ 𝑥 ∧ ∀𝑦(𝑦 ∈ 𝑥 → suc 𝑦 ∈ 𝑥)) | ||
Theorem | zfinf2 4582* | A standard version of the Axiom of Infinity, using definitions to abbreviate. Axiom Inf of [BellMachover] p. 472. (Contributed by NM, 30-Aug-1993.) |
⊢ ∃𝑥(∅ ∈ 𝑥 ∧ ∀𝑦 ∈ 𝑥 suc 𝑦 ∈ 𝑥) | ||
Syntax | com 4583 | Extend class notation to include the class of natural numbers. |
class ω | ||
Definition | df-iom 4584* |
Define the class of natural numbers as the smallest inductive set, which
is valid provided we assume the Axiom of Infinity. Definition 6.3 of
[Eisenberg] p. 82.
Note: the natural numbers ω are a subset of the ordinal numbers df-on 4362. Later, when we define complex numbers, we will be able to also define a subset of the complex numbers (df-inn 8893) with analogous properties and operations, but they will be different sets. We are unable to use the terms finite ordinal and natural number interchangeably, as shown at exmidonfin 7183. (Contributed by NM, 6-Aug-1994.) Use its alias dfom3 4585 instead for naming consistency with set.mm. (New usage is discouraged.) |
⊢ ω = ∩ {𝑥 ∣ (∅ ∈ 𝑥 ∧ ∀𝑦 ∈ 𝑥 suc 𝑦 ∈ 𝑥)} | ||
Theorem | dfom3 4585* | Alias for df-iom 4584. Use it instead of df-iom 4584 for naming consistency with set.mm. (Contributed by NM, 6-Aug-1994.) |
⊢ ω = ∩ {𝑥 ∣ (∅ ∈ 𝑥 ∧ ∀𝑦 ∈ 𝑥 suc 𝑦 ∈ 𝑥)} | ||
Theorem | omex 4586 | The existence of omega (the class of natural numbers). Axiom 7 of [TakeutiZaring] p. 43. (Contributed by NM, 6-Aug-1994.) |
⊢ ω ∈ V | ||
Theorem | peano1 4587 | Zero is a natural number. One of Peano's five postulates for arithmetic. Proposition 7.30(1) of [TakeutiZaring] p. 42. (Contributed by NM, 15-May-1994.) |
⊢ ∅ ∈ ω | ||
Theorem | peano2 4588 | The successor of any natural number is a natural number. One of Peano's five postulates for arithmetic. Proposition 7.30(2) of [TakeutiZaring] p. 42. (Contributed by NM, 3-Sep-2003.) |
⊢ (𝐴 ∈ ω → suc 𝐴 ∈ ω) | ||
Theorem | peano3 4589 | The successor of any natural number is not zero. One of Peano's five postulates for arithmetic. Proposition 7.30(3) of [TakeutiZaring] p. 42. (Contributed by NM, 3-Sep-2003.) |
⊢ (𝐴 ∈ ω → suc 𝐴 ≠ ∅) | ||
Theorem | peano4 4590 | Two natural numbers are equal iff their successors are equal, i.e. the successor function is one-to-one. One of Peano's five postulates for arithmetic. Proposition 7.30(4) of [TakeutiZaring] p. 43. (Contributed by NM, 3-Sep-2003.) |
⊢ ((𝐴 ∈ ω ∧ 𝐵 ∈ ω) → (suc 𝐴 = suc 𝐵 ↔ 𝐴 = 𝐵)) | ||
Theorem | peano5 4591* | The induction postulate: any class containing zero and closed under the successor operation contains all natural numbers. One of Peano's five postulates for arithmetic. Proposition 7.30(5) of [TakeutiZaring] p. 43. The more traditional statement of mathematical induction as a theorem schema, with a basis and an induction step, is derived from this theorem as Theorem findes 4596. (Contributed by NM, 18-Feb-2004.) |
⊢ ((∅ ∈ 𝐴 ∧ ∀𝑥 ∈ ω (𝑥 ∈ 𝐴 → suc 𝑥 ∈ 𝐴)) → ω ⊆ 𝐴) | ||
Theorem | find 4592* | The Principle of Finite Induction (mathematical induction). Corollary 7.31 of [TakeutiZaring] p. 43. The simpler hypothesis shown here was suggested in an email from "Colin" on 1-Oct-2001. The hypothesis states that 𝐴 is a set of natural numbers, zero belongs to 𝐴, and given any member of 𝐴 the member's successor also belongs to 𝐴. The conclusion is that every natural number is in 𝐴. (Contributed by NM, 22-Feb-2004.) (Proof shortened by Andrew Salmon, 27-Aug-2011.) |
⊢ (𝐴 ⊆ ω ∧ ∅ ∈ 𝐴 ∧ ∀𝑥 ∈ 𝐴 suc 𝑥 ∈ 𝐴) ⇒ ⊢ 𝐴 = ω | ||
Theorem | finds 4593* | Principle of Finite Induction (inference schema), using implicit substitutions. The first four hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. This is Metamath 100 proof #74. (Contributed by NM, 14-Apr-1995.) |
⊢ (𝑥 = ∅ → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 = suc 𝑦 → (𝜑 ↔ 𝜃)) & ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜏)) & ⊢ 𝜓 & ⊢ (𝑦 ∈ ω → (𝜒 → 𝜃)) ⇒ ⊢ (𝐴 ∈ ω → 𝜏) | ||
Theorem | finds2 4594* | Principle of Finite Induction (inference schema), using implicit substitutions. The first three hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. (Contributed by NM, 29-Nov-2002.) |
⊢ (𝑥 = ∅ → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 = suc 𝑦 → (𝜑 ↔ 𝜃)) & ⊢ (𝜏 → 𝜓) & ⊢ (𝑦 ∈ ω → (𝜏 → (𝜒 → 𝜃))) ⇒ ⊢ (𝑥 ∈ ω → (𝜏 → 𝜑)) | ||
Theorem | finds1 4595* | Principle of Finite Induction (inference schema), using implicit substitutions. The first three hypotheses establish the substitutions we need. The last two are the basis and the induction step. Theorem Schema 22 of [Suppes] p. 136. (Contributed by NM, 22-Mar-2006.) |
⊢ (𝑥 = ∅ → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 = suc 𝑦 → (𝜑 ↔ 𝜃)) & ⊢ 𝜓 & ⊢ (𝑦 ∈ ω → (𝜒 → 𝜃)) ⇒ ⊢ (𝑥 ∈ ω → 𝜑) | ||
Theorem | findes 4596 | Finite induction with explicit substitution. The first hypothesis is the basis and the second is the induction step. Theorem Schema 22 of [Suppes] p. 136. This is an alternative for Metamath 100 proof #74. (Contributed by Raph Levien, 9-Jul-2003.) |
⊢ [∅ / 𝑥]𝜑 & ⊢ (𝑥 ∈ ω → (𝜑 → [suc 𝑥 / 𝑥]𝜑)) ⇒ ⊢ (𝑥 ∈ ω → 𝜑) | ||
Theorem | nn0suc 4597* | A natural number is either 0 or a successor. Similar theorems for arbitrary sets or real numbers will not be provable (without the law of the excluded middle), but equality of natural numbers is decidable. (Contributed by NM, 27-May-1998.) |
⊢ (𝐴 ∈ ω → (𝐴 = ∅ ∨ ∃𝑥 ∈ ω 𝐴 = suc 𝑥)) | ||
Theorem | elomssom 4598 | A natural number ordinal is, as a set, included in the set of natural number ordinals. (Contributed by NM, 21-Jun-1998.) Extract this result from the previous proof of elnn 4599. (Revised by BJ, 7-Aug-2024.) |
⊢ (𝐴 ∈ ω → 𝐴 ⊆ ω) | ||
Theorem | elnn 4599 | A member of a natural number is a natural number. (Contributed by NM, 21-Jun-1998.) |
⊢ ((𝐴 ∈ 𝐵 ∧ 𝐵 ∈ ω) → 𝐴 ∈ ω) | ||
Theorem | ordom 4600 | Omega is ordinal. Theorem 7.32 of [TakeutiZaring] p. 43. (Contributed by NM, 18-Oct-1995.) |
⊢ Ord ω |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |