MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  inf3 Structured version   Visualization version   GIF version

Theorem inf3 8532
Description: Our Axiom of Infinity ax-inf 8535 implies the standard Axiom of Infinity. The hypothesis is a variant of our Axiom of Infinity provided by inf2 8520, 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 8537 and zfinf2 8539.) The main proof is provided by inf3lema 8521 through inf3lem7 8531, and this final piece eliminates the auxiliary hypothesis of inf3lem7 8531. 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 8522.)
       F_n+1 = {y<X | y^X subset F_n}  (See inf3lemc 8523.)
       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 8525.)
       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 8526.)
       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 8527.)
       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 8528.)
       Proof:  Lemmas 1 and 3.

       Lemma 5.  F_m proper_subset F_n, m < n.  (See inf3lem5 8529.)
       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 8530.)
       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 8531.)
       Q.E.D.
       
(Contributed by NM, 29-Oct-1996.)
Hypothesis
Ref Expression
inf3.1 𝑥(𝑥 ≠ ∅ ∧ 𝑥 𝑥)
Assertion
Ref Expression
inf3 ω ∈ V

Proof of Theorem inf3
Dummy variables 𝑦 𝑤 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 eqid 2622 . . 3 (𝑦 ∈ V ↦ {𝑤𝑥 ∣ (𝑤𝑥) ⊆ 𝑦}) = (𝑦 ∈ V ↦ {𝑤𝑥 ∣ (𝑤𝑥) ⊆ 𝑦})
2 eqid 2622 . . 3 (rec((𝑦 ∈ V ↦ {𝑤𝑥 ∣ (𝑤𝑥) ⊆ 𝑦}), ∅) ↾ ω) = (rec((𝑦 ∈ V ↦ {𝑤𝑥 ∣ (𝑤𝑥) ⊆ 𝑦}), ∅) ↾ ω)
3 vex 3203 . . 3 𝑥 ∈ V
41, 2, 3, 3inf3lem7 8531 . 2 ((𝑥 ≠ ∅ ∧ 𝑥 𝑥) → ω ∈ V)
5 inf3.1 . 2 𝑥(𝑥 ≠ ∅ ∧ 𝑥 𝑥)
64, 5exlimiiv 1859 1 ω ∈ V
Colors of variables: wff setvar class
Syntax hints:  wa 384  wex 1704  wcel 1990  wne 2794  {crab 2916  Vcvv 3200  cin 3573  wss 3574  c0 3915   cuni 4436  cmpt 4729  cres 5116  ωcom 7065  reccrdg 7505
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-8 1992  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-rep 4771  ax-sep 4781  ax-nul 4789  ax-pow 4843  ax-pr 4906  ax-un 6949  ax-reg 8497
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3or 1038  df-3an 1039  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-ne 2795  df-ral 2917  df-rex 2918  df-reu 2919  df-rab 2921  df-v 3202  df-sbc 3436  df-csb 3534  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-pss 3590  df-nul 3916  df-if 4087  df-pw 4160  df-sn 4178  df-pr 4180  df-tp 4182  df-op 4184  df-uni 4437  df-iun 4522  df-br 4654  df-opab 4713  df-mpt 4730  df-tr 4753  df-id 5024  df-eprel 5029  df-po 5035  df-so 5036  df-fr 5073  df-we 5075  df-xp 5120  df-rel 5121  df-cnv 5122  df-co 5123  df-dm 5124  df-rn 5125  df-res 5126  df-ima 5127  df-pred 5680  df-ord 5726  df-on 5727  df-lim 5728  df-suc 5729  df-iota 5851  df-fun 5890  df-fn 5891  df-f 5892  df-f1 5893  df-fo 5894  df-f1o 5895  df-fv 5896  df-om 7066  df-wrecs 7407  df-recs 7468  df-rdg 7506
This theorem is referenced by:  axinf2  8537
  Copyright terms: Public domain W3C validator