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

Theorem dfnn3 10911
Description: Alternate definition of the set of positive integers. Definition of positive integers in [Apostol] p. 22. (Contributed by NM, 3-Jul-2005.)
Assertion
Ref Expression
dfnn3 ℕ = {𝑥 ∣ (𝑥 ⊆ ℝ ∧ 1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}
Distinct variable group:   𝑥,𝑦

Proof of Theorem dfnn3
Dummy variable 𝑧 is distinct from all other variables.
StepHypRef Expression
1 eleq2 2677 . . . 4 (𝑥 = 𝑧 → (1 ∈ 𝑥 ↔ 1 ∈ 𝑧))
2 eleq2 2677 . . . . 5 (𝑥 = 𝑧 → ((𝑦 + 1) ∈ 𝑥 ↔ (𝑦 + 1) ∈ 𝑧))
32raleqbi1dv 3123 . . . 4 (𝑥 = 𝑧 → (∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥 ↔ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧))
41, 3anbi12d 743 . . 3 (𝑥 = 𝑧 → ((1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥) ↔ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)))
5 dfnn2 10910 . . . . 5 ℕ = {𝑧 ∣ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)}
65eqeq2i 2622 . . . 4 (𝑥 = ℕ ↔ 𝑥 = {𝑧 ∣ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)})
7 eleq2 2677 . . . . 5 (𝑥 = ℕ → (1 ∈ 𝑥 ↔ 1 ∈ ℕ))
8 eleq2 2677 . . . . . 6 (𝑥 = ℕ → ((𝑦 + 1) ∈ 𝑥 ↔ (𝑦 + 1) ∈ ℕ))
98raleqbi1dv 3123 . . . . 5 (𝑥 = ℕ → (∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥 ↔ ∀𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ))
107, 9anbi12d 743 . . . 4 (𝑥 = ℕ → ((1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥) ↔ (1 ∈ ℕ ∧ ∀𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ)))
116, 10sylbir 224 . . 3 (𝑥 = {𝑧 ∣ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)} → ((1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥) ↔ (1 ∈ ℕ ∧ ∀𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ)))
12 nnssre 10901 . . . . 5 ℕ ⊆ ℝ
135, 12eqsstr3i 3599 . . . 4 {𝑧 ∣ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)} ⊆ ℝ
14 1nn 10908 . . . . 5 1 ∈ ℕ
15 peano2nn 10909 . . . . . 6 (𝑦 ∈ ℕ → (𝑦 + 1) ∈ ℕ)
1615rgen 2906 . . . . 5 𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ
1714, 16pm3.2i 470 . . . 4 (1 ∈ ℕ ∧ ∀𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ)
1813, 17pm3.2i 470 . . 3 ( {𝑧 ∣ (1 ∈ 𝑧 ∧ ∀𝑦𝑧 (𝑦 + 1) ∈ 𝑧)} ⊆ ℝ ∧ (1 ∈ ℕ ∧ ∀𝑦 ∈ ℕ (𝑦 + 1) ∈ ℕ))
194, 11, 18intabs 4752 . 2 {𝑥 ∣ (𝑥 ⊆ ℝ ∧ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥))} = {𝑥 ∣ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}
20 3anass 1035 . . . 4 ((𝑥 ⊆ ℝ ∧ 1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥) ↔ (𝑥 ⊆ ℝ ∧ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)))
2120abbii 2726 . . 3 {𝑥 ∣ (𝑥 ⊆ ℝ ∧ 1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)} = {𝑥 ∣ (𝑥 ⊆ ℝ ∧ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥))}
2221inteqi 4414 . 2 {𝑥 ∣ (𝑥 ⊆ ℝ ∧ 1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)} = {𝑥 ∣ (𝑥 ⊆ ℝ ∧ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥))}
23 dfnn2 10910 . 2 ℕ = {𝑥 ∣ (1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}
2419, 22, 233eqtr4ri 2643 1 ℕ = {𝑥 ∣ (𝑥 ⊆ ℝ ∧ 1 ∈ 𝑥 ∧ ∀𝑦𝑥 (𝑦 + 1) ∈ 𝑥)}
Colors of variables: wff setvar class
Syntax hints:  wb 195  wa 383  w3a 1031   = wceq 1475  wcel 1977  {cab 2596  wral 2896  wss 3540   cint 4410  (class class class)co 6549  cr 9814  1c1 9816   + caddc 9818  cn 10897
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1713  ax-4 1728  ax-5 1827  ax-6 1875  ax-7 1922  ax-8 1979  ax-9 1986  ax-10 2006  ax-11 2021  ax-12 2034  ax-13 2234  ax-ext 2590  ax-sep 4709  ax-nul 4717  ax-pow 4769  ax-pr 4833  ax-un 6847  ax-cnex 9871  ax-resscn 9872  ax-1cn 9873  ax-icn 9874  ax-addcl 9875  ax-addrcl 9876  ax-mulcl 9877  ax-mulrcl 9878  ax-i2m1 9883  ax-1ne0 9884  ax-rrecex 9887  ax-cnre 9888
This theorem depends on definitions:  df-bi 196  df-or 384  df-an 385  df-3or 1032  df-3an 1033  df-tru 1478  df-ex 1696  df-nf 1701  df-sb 1868  df-eu 2462  df-mo 2463  df-clab 2597  df-cleq 2603  df-clel 2606  df-nfc 2740  df-ne 2782  df-ral 2901  df-rex 2902  df-reu 2903  df-rab 2905  df-v 3175  df-sbc 3403  df-csb 3500  df-dif 3543  df-un 3545  df-in 3547  df-ss 3554  df-pss 3556  df-nul 3875  df-if 4037  df-pw 4110  df-sn 4126  df-pr 4128  df-tp 4130  df-op 4132  df-uni 4373  df-int 4411  df-iun 4457  df-br 4584  df-opab 4644  df-mpt 4645  df-tr 4681  df-eprel 4949  df-id 4953  df-po 4959  df-so 4960  df-fr 4997  df-we 4999  df-xp 5044  df-rel 5045  df-cnv 5046  df-co 5047  df-dm 5048  df-rn 5049  df-res 5050  df-ima 5051  df-pred 5597  df-ord 5643  df-on 5644  df-lim 5645  df-suc 5646  df-iota 5768  df-fun 5806  df-fn 5807  df-f 5808  df-f1 5809  df-fo 5810  df-f1o 5811  df-fv 5812  df-ov 6552  df-om 6958  df-wrecs 7294  df-recs 7355  df-rdg 7393  df-nn 10898
This theorem is referenced by: (None)
  Copyright terms: Public domain W3C validator