| Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > ILE Home > Th. List > finds | GIF version | ||
| Description: 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.) |
| Ref | Expression |
|---|---|
| finds.1 | ⊢ (𝑥 = ∅ → (𝜑 ↔ 𝜓)) |
| finds.2 | ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜒)) |
| finds.3 | ⊢ (𝑥 = suc 𝑦 → (𝜑 ↔ 𝜃)) |
| finds.4 | ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜏)) |
| finds.5 | ⊢ 𝜓 |
| finds.6 | ⊢ (𝑦 ∈ ω → (𝜒 → 𝜃)) |
| Ref | Expression |
|---|---|
| finds | ⊢ (𝐴 ∈ ω → 𝜏) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | finds.5 | . . . . 5 ⊢ 𝜓 | |
| 2 | 0ex 4161 | . . . . . 6 ⊢ ∅ ∈ V | |
| 3 | finds.1 | . . . . . 6 ⊢ (𝑥 = ∅ → (𝜑 ↔ 𝜓)) | |
| 4 | 2, 3 | elab 2908 | . . . . 5 ⊢ (∅ ∈ {𝑥 ∣ 𝜑} ↔ 𝜓) |
| 5 | 1, 4 | mpbir 146 | . . . 4 ⊢ ∅ ∈ {𝑥 ∣ 𝜑} |
| 6 | finds.6 | . . . . . 6 ⊢ (𝑦 ∈ ω → (𝜒 → 𝜃)) | |
| 7 | vex 2766 | . . . . . . 7 ⊢ 𝑦 ∈ V | |
| 8 | finds.2 | . . . . . . 7 ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜒)) | |
| 9 | 7, 8 | elab 2908 | . . . . . 6 ⊢ (𝑦 ∈ {𝑥 ∣ 𝜑} ↔ 𝜒) |
| 10 | 7 | sucex 4536 | . . . . . . 7 ⊢ suc 𝑦 ∈ V |
| 11 | finds.3 | . . . . . . 7 ⊢ (𝑥 = suc 𝑦 → (𝜑 ↔ 𝜃)) | |
| 12 | 10, 11 | elab 2908 | . . . . . 6 ⊢ (suc 𝑦 ∈ {𝑥 ∣ 𝜑} ↔ 𝜃) |
| 13 | 6, 9, 12 | 3imtr4g 205 | . . . . 5 ⊢ (𝑦 ∈ ω → (𝑦 ∈ {𝑥 ∣ 𝜑} → suc 𝑦 ∈ {𝑥 ∣ 𝜑})) |
| 14 | 13 | rgen 2550 | . . . 4 ⊢ ∀𝑦 ∈ ω (𝑦 ∈ {𝑥 ∣ 𝜑} → suc 𝑦 ∈ {𝑥 ∣ 𝜑}) |
| 15 | peano5 4635 | . . . 4 ⊢ ((∅ ∈ {𝑥 ∣ 𝜑} ∧ ∀𝑦 ∈ ω (𝑦 ∈ {𝑥 ∣ 𝜑} → suc 𝑦 ∈ {𝑥 ∣ 𝜑})) → ω ⊆ {𝑥 ∣ 𝜑}) | |
| 16 | 5, 14, 15 | mp2an 426 | . . 3 ⊢ ω ⊆ {𝑥 ∣ 𝜑} |
| 17 | 16 | sseli 3180 | . 2 ⊢ (𝐴 ∈ ω → 𝐴 ∈ {𝑥 ∣ 𝜑}) |
| 18 | finds.4 | . . 3 ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜏)) | |
| 19 | 18 | elabg 2910 | . 2 ⊢ (𝐴 ∈ ω → (𝐴 ∈ {𝑥 ∣ 𝜑} ↔ 𝜏)) |
| 20 | 17, 19 | mpbid 147 | 1 ⊢ (𝐴 ∈ ω → 𝜏) |
| Colors of variables: wff set class |
| Syntax hints: → wi 4 ↔ wb 105 = wceq 1364 ∈ wcel 2167 {cab 2182 ∀wral 2475 ⊆ wss 3157 ∅c0 3451 suc csuc 4401 ωcom 4627 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-in1 615 ax-in2 616 ax-io 710 ax-5 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-13 2169 ax-14 2170 ax-ext 2178 ax-sep 4152 ax-nul 4160 ax-pow 4208 ax-pr 4243 ax-un 4469 ax-iinf 4625 |
| This theorem depends on definitions: df-bi 117 df-3an 982 df-tru 1367 df-nf 1475 df-sb 1777 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ral 2480 df-rex 2481 df-v 2765 df-dif 3159 df-un 3161 df-in 3163 df-ss 3170 df-nul 3452 df-pw 3608 df-sn 3629 df-pr 3630 df-uni 3841 df-int 3876 df-suc 4407 df-iom 4628 |
| This theorem is referenced by: findes 4640 nn0suc 4641 elomssom 4642 ordom 4644 nndceq0 4655 0elnn 4656 omsinds 4659 nna0r 6537 nnm0r 6538 nnsucelsuc 6550 nneneq 6919 php5 6920 php5dom 6925 fidcenumlemrk 7021 fidcenumlemr 7022 nninfninc 7190 nnnninfeq 7195 nnnninfeq2 7196 frec2uzltd 10497 frecuzrdgg 10510 seq3val 10554 seqvalcd 10555 omgadd 10896 zfz1iso 10935 ennnfonelemhom 12642 nninfsellemdc 15664 |
| Copyright terms: Public domain | W3C validator |