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

Theorem finds2 6986
Description: 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.)
Hypotheses
Ref Expression
finds2.1 (𝑥 = ∅ → (𝜑𝜓))
finds2.2 (𝑥 = 𝑦 → (𝜑𝜒))
finds2.3 (𝑥 = suc 𝑦 → (𝜑𝜃))
finds2.4 (𝜏𝜓)
finds2.5 (𝑦 ∈ ω → (𝜏 → (𝜒𝜃)))
Assertion
Ref Expression
finds2 (𝑥 ∈ ω → (𝜏𝜑))
Distinct variable groups:   𝑥,𝑦,𝜏   𝜓,𝑥   𝜒,𝑥   𝜃,𝑥   𝜑,𝑦
Allowed substitution hints:   𝜑(𝑥)   𝜓(𝑦)   𝜒(𝑦)   𝜃(𝑦)

Proof of Theorem finds2
StepHypRef Expression
1 finds2.4 . . . . 5 (𝜏𝜓)
2 0ex 4718 . . . . . 6 ∅ ∈ V
3 finds2.1 . . . . . . 7 (𝑥 = ∅ → (𝜑𝜓))
43imbi2d 329 . . . . . 6 (𝑥 = ∅ → ((𝜏𝜑) ↔ (𝜏𝜓)))
52, 4elab 3319 . . . . 5 (∅ ∈ {𝑥 ∣ (𝜏𝜑)} ↔ (𝜏𝜓))
61, 5mpbir 220 . . . 4 ∅ ∈ {𝑥 ∣ (𝜏𝜑)}
7 finds2.5 . . . . . . 7 (𝑦 ∈ ω → (𝜏 → (𝜒𝜃)))
87a2d 29 . . . . . 6 (𝑦 ∈ ω → ((𝜏𝜒) → (𝜏𝜃)))
9 vex 3176 . . . . . . 7 𝑦 ∈ V
10 finds2.2 . . . . . . . 8 (𝑥 = 𝑦 → (𝜑𝜒))
1110imbi2d 329 . . . . . . 7 (𝑥 = 𝑦 → ((𝜏𝜑) ↔ (𝜏𝜒)))
129, 11elab 3319 . . . . . 6 (𝑦 ∈ {𝑥 ∣ (𝜏𝜑)} ↔ (𝜏𝜒))
139sucex 6903 . . . . . . 7 suc 𝑦 ∈ V
14 finds2.3 . . . . . . . 8 (𝑥 = suc 𝑦 → (𝜑𝜃))
1514imbi2d 329 . . . . . . 7 (𝑥 = suc 𝑦 → ((𝜏𝜑) ↔ (𝜏𝜃)))
1613, 15elab 3319 . . . . . 6 (suc 𝑦 ∈ {𝑥 ∣ (𝜏𝜑)} ↔ (𝜏𝜃))
178, 12, 163imtr4g 284 . . . . 5 (𝑦 ∈ ω → (𝑦 ∈ {𝑥 ∣ (𝜏𝜑)} → suc 𝑦 ∈ {𝑥 ∣ (𝜏𝜑)}))
1817rgen 2906 . . . 4 𝑦 ∈ ω (𝑦 ∈ {𝑥 ∣ (𝜏𝜑)} → suc 𝑦 ∈ {𝑥 ∣ (𝜏𝜑)})
19 peano5 6981 . . . 4 ((∅ ∈ {𝑥 ∣ (𝜏𝜑)} ∧ ∀𝑦 ∈ ω (𝑦 ∈ {𝑥 ∣ (𝜏𝜑)} → suc 𝑦 ∈ {𝑥 ∣ (𝜏𝜑)})) → ω ⊆ {𝑥 ∣ (𝜏𝜑)})
206, 18, 19mp2an 704 . . 3 ω ⊆ {𝑥 ∣ (𝜏𝜑)}
2120sseli 3564 . 2 (𝑥 ∈ ω → 𝑥 ∈ {𝑥 ∣ (𝜏𝜑)})
22 abid 2598 . 2 (𝑥 ∈ {𝑥 ∣ (𝜏𝜑)} ↔ (𝜏𝜑))
2321, 22sylib 207 1 (𝑥 ∈ ω → (𝜏𝜑))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 195   = wceq 1475  wcel 1977  {cab 2596  wral 2896  wss 3540  c0 3874  suc csuc 5642  ωcom 6957
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-pr 4833  ax-un 6847
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-rab 2905  df-v 3175  df-sbc 3403  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-br 4584  df-opab 4644  df-tr 4681  df-eprel 4949  df-po 4959  df-so 4960  df-fr 4997  df-we 4999  df-ord 5643  df-on 5644  df-lim 5645  df-suc 5646  df-om 6958
This theorem is referenced by:  finds1  6987  onnseq  7328  nnacl  7578  nnmcl  7579  nnecl  7580  nnacom  7584  nnaass  7589  nndi  7590  nnmass  7591  nnmsucr  7592  nnmcom  7593  nnmordi  7598  omsmolem  7620  isinf  8058  unblem2  8098  fiint  8122  dffi3  8220  card2inf  8343  cantnfle  8451  cantnflt  8452  cantnflem1  8469  cnfcom  8480  trcl  8487  fseqenlem1  8730  infpssrlem3  9010  fin23lem26  9030  axdc3lem2  9156  axdc4lem  9160  axdclem2  9225  wunr1om  9420  wuncval2  9448  tskr1om  9468  grothomex  9530  peano5nni  10900  neibastop2lem  31525  finxpreclem6  32409
  Copyright terms: Public domain W3C validator