ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  nnind GIF version

Theorem nnind 8055
Description: Principle of Mathematical Induction (inference schema). The first four hypotheses give us the substitution instances we need; the last two are the basis and the induction step. See nnaddcl 8059 for an example of its use. This is an alternative for Metamath 100 proof #74. (Contributed by NM, 10-Jan-1997.) (Revised by Mario Carneiro, 16-Jun-2013.)
Hypotheses
Ref Expression
nnind.1 (𝑥 = 1 → (𝜑𝜓))
nnind.2 (𝑥 = 𝑦 → (𝜑𝜒))
nnind.3 (𝑥 = (𝑦 + 1) → (𝜑𝜃))
nnind.4 (𝑥 = 𝐴 → (𝜑𝜏))
nnind.5 𝜓
nnind.6 (𝑦 ∈ ℕ → (𝜒𝜃))
Assertion
Ref Expression
nnind (𝐴 ∈ ℕ → 𝜏)
Distinct variable groups:   𝑥,𝑦   𝑥,𝐴   𝜓,𝑥   𝜒,𝑥   𝜃,𝑥   𝜏,𝑥   𝜑,𝑦
Allowed substitution hints:   𝜑(𝑥)   𝜓(𝑦)   𝜒(𝑦)   𝜃(𝑦)   𝜏(𝑦)   𝐴(𝑦)

Proof of Theorem nnind
StepHypRef Expression
1 1nn 8050 . . . . . 6 1 ∈ ℕ
2 nnind.5 . . . . . 6 𝜓
3 nnind.1 . . . . . . 7 (𝑥 = 1 → (𝜑𝜓))
43elrab 2749 . . . . . 6 (1 ∈ {𝑥 ∈ ℕ ∣ 𝜑} ↔ (1 ∈ ℕ ∧ 𝜓))
51, 2, 4mpbir2an 883 . . . . 5 1 ∈ {𝑥 ∈ ℕ ∣ 𝜑}
6 elrabi 2746 . . . . . . 7 (𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} → 𝑦 ∈ ℕ)
7 peano2nn 8051 . . . . . . . . . 10 (𝑦 ∈ ℕ → (𝑦 + 1) ∈ ℕ)
87a1d 22 . . . . . . . . 9 (𝑦 ∈ ℕ → (𝑦 ∈ ℕ → (𝑦 + 1) ∈ ℕ))
9 nnind.6 . . . . . . . . 9 (𝑦 ∈ ℕ → (𝜒𝜃))
108, 9anim12d 328 . . . . . . . 8 (𝑦 ∈ ℕ → ((𝑦 ∈ ℕ ∧ 𝜒) → ((𝑦 + 1) ∈ ℕ ∧ 𝜃)))
11 nnind.2 . . . . . . . . 9 (𝑥 = 𝑦 → (𝜑𝜒))
1211elrab 2749 . . . . . . . 8 (𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} ↔ (𝑦 ∈ ℕ ∧ 𝜒))
13 nnind.3 . . . . . . . . 9 (𝑥 = (𝑦 + 1) → (𝜑𝜃))
1413elrab 2749 . . . . . . . 8 ((𝑦 + 1) ∈ {𝑥 ∈ ℕ ∣ 𝜑} ↔ ((𝑦 + 1) ∈ ℕ ∧ 𝜃))
1510, 12, 143imtr4g 203 . . . . . . 7 (𝑦 ∈ ℕ → (𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} → (𝑦 + 1) ∈ {𝑥 ∈ ℕ ∣ 𝜑}))
166, 15mpcom 36 . . . . . 6 (𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} → (𝑦 + 1) ∈ {𝑥 ∈ ℕ ∣ 𝜑})
1716rgen 2416 . . . . 5 𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} (𝑦 + 1) ∈ {𝑥 ∈ ℕ ∣ 𝜑}
18 peano5nni 8042 . . . . 5 ((1 ∈ {𝑥 ∈ ℕ ∣ 𝜑} ∧ ∀𝑦 ∈ {𝑥 ∈ ℕ ∣ 𝜑} (𝑦 + 1) ∈ {𝑥 ∈ ℕ ∣ 𝜑}) → ℕ ⊆ {𝑥 ∈ ℕ ∣ 𝜑})
195, 17, 18mp2an 416 . . . 4 ℕ ⊆ {𝑥 ∈ ℕ ∣ 𝜑}
2019sseli 2995 . . 3 (𝐴 ∈ ℕ → 𝐴 ∈ {𝑥 ∈ ℕ ∣ 𝜑})
21 nnind.4 . . . 4 (𝑥 = 𝐴 → (𝜑𝜏))
2221elrab 2749 . . 3 (𝐴 ∈ {𝑥 ∈ ℕ ∣ 𝜑} ↔ (𝐴 ∈ ℕ ∧ 𝜏))
2320, 22sylib 120 . 2 (𝐴 ∈ ℕ → (𝐴 ∈ ℕ ∧ 𝜏))
2423simprd 112 1 (𝐴 ∈ ℕ → 𝜏)
Colors of variables: wff set class
Syntax hints:  wi 4  wa 102  wb 103   = wceq 1284  wcel 1433  wral 2348  {crab 2352  wss 2973  (class class class)co 5532  1c1 6982   + caddc 6984  cn 8039
This theorem was proved from axioms:  ax-1 5  ax-2 6  ax-mp 7  ax-ia1 104  ax-ia2 105  ax-ia3 106  ax-io 662  ax-5 1376  ax-7 1377  ax-gen 1378  ax-ie1 1422  ax-ie2 1423  ax-8 1435  ax-10 1436  ax-11 1437  ax-i12 1438  ax-bndl 1439  ax-4 1440  ax-17 1459  ax-i9 1463  ax-ial 1467  ax-i5r 1468  ax-ext 2063  ax-sep 3896  ax-cnex 7067  ax-resscn 7068  ax-1re 7070  ax-addrcl 7073
This theorem depends on definitions:  df-bi 115  df-3an 921  df-tru 1287  df-nf 1390  df-sb 1686  df-clab 2068  df-cleq 2074  df-clel 2077  df-nfc 2208  df-ral 2353  df-rex 2354  df-rab 2357  df-v 2603  df-un 2977  df-in 2979  df-ss 2986  df-sn 3404  df-pr 3405  df-op 3407  df-uni 3602  df-int 3637  df-br 3786  df-iota 4887  df-fv 4930  df-ov 5535  df-inn 8040
This theorem is referenced by:  nnindALT  8056  nn1m1nn  8057  nnaddcl  8059  nnmulcl  8060  nnge1  8062  nn1gt1  8072  nnsub  8077  zaddcllempos  8388  zaddcllemneg  8390  nneoor  8449  peano5uzti  8455  nn0ind-raph  8464  indstr  8681  qbtwnzlemshrink  9258  expivallem  9477  expcllem  9487  expap0  9506  resqrexlemover  9896  resqrexlemlo  9899  resqrexlemcalc3  9902  gcdmultiple  10409  rplpwr  10416  prmind2  10502  prmdvdsexp  10527  sqrt2irr  10541  pw2dvdslemn  10543
  Copyright terms: Public domain W3C validator