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

Theorem nnaddcl 7934
Description: Closure of addition of positive integers, proved by induction on the second addend. (Contributed by NM, 12-Jan-1997.)
Assertion
Ref Expression
nnaddcl ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 + 𝐵) ∈ ℕ)

Proof of Theorem nnaddcl
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 oveq2 5520 . . . . 5 (𝑥 = 1 → (𝐴 + 𝑥) = (𝐴 + 1))
21eleq1d 2106 . . . 4 (𝑥 = 1 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 1) ∈ ℕ))
32imbi2d 219 . . 3 (𝑥 = 1 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 1) ∈ ℕ)))
4 oveq2 5520 . . . . 5 (𝑥 = 𝑦 → (𝐴 + 𝑥) = (𝐴 + 𝑦))
54eleq1d 2106 . . . 4 (𝑥 = 𝑦 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 𝑦) ∈ ℕ))
65imbi2d 219 . . 3 (𝑥 = 𝑦 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 𝑦) ∈ ℕ)))
7 oveq2 5520 . . . . 5 (𝑥 = (𝑦 + 1) → (𝐴 + 𝑥) = (𝐴 + (𝑦 + 1)))
87eleq1d 2106 . . . 4 (𝑥 = (𝑦 + 1) → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + (𝑦 + 1)) ∈ ℕ))
98imbi2d 219 . . 3 (𝑥 = (𝑦 + 1) → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ)))
10 oveq2 5520 . . . . 5 (𝑥 = 𝐵 → (𝐴 + 𝑥) = (𝐴 + 𝐵))
1110eleq1d 2106 . . . 4 (𝑥 = 𝐵 → ((𝐴 + 𝑥) ∈ ℕ ↔ (𝐴 + 𝐵) ∈ ℕ))
1211imbi2d 219 . . 3 (𝑥 = 𝐵 → ((𝐴 ∈ ℕ → (𝐴 + 𝑥) ∈ ℕ) ↔ (𝐴 ∈ ℕ → (𝐴 + 𝐵) ∈ ℕ)))
13 peano2nn 7926 . . 3 (𝐴 ∈ ℕ → (𝐴 + 1) ∈ ℕ)
14 peano2nn 7926 . . . . . 6 ((𝐴 + 𝑦) ∈ ℕ → ((𝐴 + 𝑦) + 1) ∈ ℕ)
15 nncn 7922 . . . . . . . 8 (𝐴 ∈ ℕ → 𝐴 ∈ ℂ)
16 nncn 7922 . . . . . . . 8 (𝑦 ∈ ℕ → 𝑦 ∈ ℂ)
17 ax-1cn 6977 . . . . . . . . 9 1 ∈ ℂ
18 addass 7011 . . . . . . . . 9 ((𝐴 ∈ ℂ ∧ 𝑦 ∈ ℂ ∧ 1 ∈ ℂ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1)))
1917, 18mp3an3 1221 . . . . . . . 8 ((𝐴 ∈ ℂ ∧ 𝑦 ∈ ℂ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1)))
2015, 16, 19syl2an 273 . . . . . . 7 ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → ((𝐴 + 𝑦) + 1) = (𝐴 + (𝑦 + 1)))
2120eleq1d 2106 . . . . . 6 ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → (((𝐴 + 𝑦) + 1) ∈ ℕ ↔ (𝐴 + (𝑦 + 1)) ∈ ℕ))
2214, 21syl5ib 143 . . . . 5 ((𝐴 ∈ ℕ ∧ 𝑦 ∈ ℕ) → ((𝐴 + 𝑦) ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ))
2322expcom 109 . . . 4 (𝑦 ∈ ℕ → (𝐴 ∈ ℕ → ((𝐴 + 𝑦) ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ)))
2423a2d 23 . . 3 (𝑦 ∈ ℕ → ((𝐴 ∈ ℕ → (𝐴 + 𝑦) ∈ ℕ) → (𝐴 ∈ ℕ → (𝐴 + (𝑦 + 1)) ∈ ℕ)))
253, 6, 9, 12, 13, 24nnind 7930 . 2 (𝐵 ∈ ℕ → (𝐴 ∈ ℕ → (𝐴 + 𝐵) ∈ ℕ))
2625impcom 116 1 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → (𝐴 + 𝐵) ∈ ℕ)
Colors of variables: wff set class
Syntax hints:  wi 4  wa 97   = wceq 1243  wcel 1393  (class class class)co 5512  cc 6887  1c1 6890   + caddc 6892  cn 7914
This theorem was proved from axioms:  ax-1 5  ax-2 6  ax-mp 7  ax-ia1 99  ax-ia2 100  ax-ia3 101  ax-io 630  ax-5 1336  ax-7 1337  ax-gen 1338  ax-ie1 1382  ax-ie2 1383  ax-8 1395  ax-10 1396  ax-11 1397  ax-i12 1398  ax-bndl 1399  ax-4 1400  ax-17 1419  ax-i9 1423  ax-ial 1427  ax-i5r 1428  ax-ext 2022  ax-sep 3875  ax-cnex 6975  ax-resscn 6976  ax-1cn 6977  ax-1re 6978  ax-addrcl 6981  ax-addass 6986
This theorem depends on definitions:  df-bi 110  df-3an 887  df-tru 1246  df-nf 1350  df-sb 1646  df-clab 2027  df-cleq 2033  df-clel 2036  df-nfc 2167  df-ral 2311  df-rex 2312  df-rab 2315  df-v 2559  df-un 2922  df-in 2924  df-ss 2931  df-sn 3381  df-pr 3382  df-op 3384  df-uni 3581  df-int 3616  df-br 3765  df-iota 4867  df-fv 4910  df-ov 5515  df-inn 7915
This theorem is referenced by:  nnmulcl  7935  nn2ge  7946  nnaddcld  7961  nnnn0addcl  8212  nn0addcl  8217
  Copyright terms: Public domain W3C validator