![]() |
Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > MPE Home > Th. List > nnaddcl | Structured version Visualization version GIF version |
Description: Closure of addition of positive integers, proved by induction on the second addend. (Contributed by NM, 12-Jan-1997.) |
Ref | Expression |
---|---|
nnaddcl | ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 + 𝐵) ∈ ℕ) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | oveq2 7031 | . . . . 5 ⊢ (𝑥 = 1 → (𝐴 + 𝑥) = (𝐴 + 1)) | |
2 | 1 | eleq1d 2869 | . . . 4 ⊢ (𝑥 = 1 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 1) ∈ ℕ)) |
3 | 2 | imbi2d 342 | . . 3 ⊢ (𝑥 = 1 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 1) ∈ ℕ))) |
4 | oveq2 7031 | . . . . 5 ⊢ (𝑥 = 𝑦 → (𝐴 + 𝑥) = (𝐴 + 𝑦)) | |
5 | 4 | eleq1d 2869 | . . . 4 ⊢ (𝑥 = 𝑦 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 𝑦) ∈ ℕ)) |
6 | 5 | imbi2d 342 | . . 3 ⊢ (𝑥 = 𝑦 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 𝑦) ∈ ℕ))) |
7 | oveq2 7031 | . . . . 5 ⊢ (𝑥 = (𝑦 + 1) → (𝐴 + 𝑥) = (𝐴 + (𝑦 + 1))) | |
8 | 7 | eleq1d 2869 | . . . 4 ⊢ (𝑥 = (𝑦 + 1) → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + (𝑦 + 1)) ∈ ℕ)) |
9 | 8 | imbi2d 342 | . . 3 ⊢ (𝑥 = (𝑦 + 1) → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ))) |
10 | oveq2 7031 | . . . . 5 ⊢ (𝑥 = 𝐵 → (𝐴 + 𝑥) = (𝐴 + 𝐵)) | |
11 | 10 | eleq1d 2869 | . . . 4 ⊢ (𝑥 = 𝐵 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 𝐵) ∈ ℕ)) |
12 | 11 | imbi2d 342 | . . 3 ⊢ (𝑥 = 𝐵 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 𝐵) ∈ ℕ))) |
13 | peano2nn 11504 | . . 3 ⊢ (𝐴 ∈ ℕ → (𝐴 + 1) ∈ ℕ) | |
14 | peano2nn 11504 | . . . . . 6 ⊢ ((𝐴 + 𝑦) ∈ ℕ → ((𝐴 + 𝑦) + 1) ∈ ℕ) | |
15 | nncn 11500 | . . . . . . . 8 ⊢ (𝐴 ∈ ℕ → 𝐴 ∈ ℂ) | |
16 | nncn 11500 | . . . . . . . 8 ⊢ (𝑦 ∈ ℕ → 𝑦 ∈ ℂ) | |
17 | ax-1cn 10448 | . . . . . . . . 9 ⊢ 1 ∈ ℂ | |
18 | addass 10477 | . . . . . . . . 9 ⊢ ((𝐴 ∈ ℂ ∧ 𝑦 ∈ ℂ ∧ 1 ∈ ℂ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1))) | |
19 | 17, 18 | mp3an3 1442 | . . . . . . . 8 ⊢ ((𝐴 ∈ ℂ ∧ 𝑦 ∈ ℂ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1))) |
20 | 15, 16, 19 | syl2an 595 | . . . . . . 7 ⊢ ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1))) |
21 | 20 | eleq1d 2869 | . . . . . 6 ⊢ ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → (((𝐴 + 𝑦) + 1) ∈ ℕ ↔ (𝐴 + (𝑦 + 1)) ∈ ℕ)) |
22 | 14, 21 | syl5ib 245 | . . . . 5 ⊢ ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → ((𝐴 + 𝑦) ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ)) |
23 | 22 | expcom 414 | . . . 4 ⊢ (𝑦 ∈ ℕ → (𝐴 ∈ ℕ → ((𝐴 + 𝑦) ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ))) |
24 | 23 | a2d 29 | . . 3 ⊢ (𝑦 ∈ ℕ → ((𝐴 ∈ ℕ → (𝐴 + 𝑦) ∈ ℕ) → (𝐴 ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ))) |
25 | 3, 6, 9, 12, 13, 24 | nnind 11510 | . 2 ⊢ (𝐵 ∈ ℕ → (𝐴 ∈ ℕ → (𝐴 + 𝐵) ∈ ℕ)) |
26 | 25 | impcom 408 | 1 ⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 + 𝐵) ∈ ℕ) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 396 = wceq 1525 ∈ wcel 2083 (class class class)co 7023 ℂcc 10388 1c1 10391 + caddc 10393 ℕcn 11492 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1781 ax-4 1795 ax-5 1892 ax-6 1951 ax-7 1996 ax-8 2085 ax-9 2093 ax-10 2114 ax-11 2128 ax-12 2143 ax-13 2346 ax-ext 2771 ax-sep 5101 ax-nul 5108 ax-pow 5164 ax-pr 5228 ax-un 7326 ax-1cn 10448 ax-addcl 10450 ax-addass 10455 |
This theorem depends on definitions: df-bi 208 df-an 397 df-or 843 df-3or 1081 df-3an 1082 df-tru 1528 df-ex 1766 df-nf 1770 df-sb 2045 df-mo 2578 df-eu 2614 df-clab 2778 df-cleq 2790 df-clel 2865 df-nfc 2937 df-ne 2987 df-ral 3112 df-rex 3113 df-reu 3114 df-rab 3116 df-v 3442 df-sbc 3712 df-csb 3818 df-dif 3868 df-un 3870 df-in 3872 df-ss 3880 df-pss 3882 df-nul 4218 df-if 4388 df-pw 4461 df-sn 4479 df-pr 4481 df-tp 4483 df-op 4485 df-uni 4752 df-iun 4833 df-br 4969 df-opab 5031 df-mpt 5048 df-tr 5071 df-id 5355 df-eprel 5360 df-po 5369 df-so 5370 df-fr 5409 df-we 5411 df-xp 5456 df-rel 5457 df-cnv 5458 df-co 5459 df-dm 5460 df-rn 5461 df-res 5462 df-ima 5463 df-pred 6030 df-ord 6076 df-on 6077 df-lim 6078 df-suc 6079 df-iota 6196 df-fun 6234 df-fn 6235 df-f 6236 df-f1 6237 df-fo 6238 df-f1o 6239 df-fv 6240 df-ov 7026 df-om 7444 df-wrecs 7805 df-recs 7867 df-rdg 7905 df-nn 11493 |
This theorem is referenced by: nnmulcl 11515 nnaddcld 11543 nnnn0addcl 11781 nn0addcl 11786 zaddcl 11876 9p1e10 11954 pythagtriplem4 15989 vdwapun 16143 vdwap1 16146 vdwlem2 16151 prmgaplem7 16226 prmgapprmolem 16230 mulgnndir 18014 uniioombllem3 23873 ballotlem1 31357 ballotlem2 31359 ballotlemfmpn 31365 ballotlem4 31369 ballotlemimin 31376 ballotlemsdom 31382 ballotlemsel1i 31383 ballotlemfrceq 31399 ballotlemfrcn0 31400 ballotlem1ri 31405 ballotth 31408 nndivsub 33416 gbepos 43427 gbowpos 43428 nnsgrpmgm 43587 |
Copyright terms: Public domain | W3C validator |