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

Theorem sadadd 15189
Description: For sequences that correspond to valid integers, the adder sequence function produces the sequence for the sum. This is effectively a proof of the correctness of the ripple carry adder, implemented with logic gates corresponding to df-had 1533 and df-cad 1546.

It is interesting to consider in what sense the sadd function can be said to be "adding" things outside the range of the bits function, that is, when adding sequences that are not eventually constant and so do not denote any integer. The correct interpretation is that the sequences are representations of 2-adic integers, which have a natural ring structure. (Contributed by Mario Carneiro, 9-Sep-2016.)

Assertion
Ref Expression
sadadd ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((bits‘𝐴) sadd (bits‘𝐵)) = (bits‘(𝐴 + 𝐵)))

Proof of Theorem sadadd
Dummy variables 𝑘 𝑐 𝑚 𝑛 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 bitsss 15148 . . . . . 6 (bits‘𝐴) ⊆ ℕ0
2 bitsss 15148 . . . . . 6 (bits‘𝐵) ⊆ ℕ0
3 sadcl 15184 . . . . . 6 (((bits‘𝐴) ⊆ ℕ0 ∧ (bits‘𝐵) ⊆ ℕ0) → ((bits‘𝐴) sadd (bits‘𝐵)) ⊆ ℕ0)
41, 2, 3mp2an 708 . . . . 5 ((bits‘𝐴) sadd (bits‘𝐵)) ⊆ ℕ0
54sseli 3599 . . . 4 (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) → 𝑘 ∈ ℕ0)
65a1i 11 . . 3 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) → 𝑘 ∈ ℕ0))
7 bitsss 15148 . . . . 5 (bits‘(𝐴 + 𝐵)) ⊆ ℕ0
87sseli 3599 . . . 4 (𝑘 ∈ (bits‘(𝐴 + 𝐵)) → 𝑘 ∈ ℕ0)
98a1i 11 . . 3 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑘 ∈ (bits‘(𝐴 + 𝐵)) → 𝑘 ∈ ℕ0))
10 eqid 2622 . . . . . . . . 9 seq0((𝑐 ∈ 2𝑜, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚 ∈ (bits‘𝐴), 𝑚 ∈ (bits‘𝐵), ∅ ∈ 𝑐), 1𝑜, ∅)), (𝑛 ∈ ℕ0 ↦ if(𝑛 = 0, ∅, (𝑛 − 1)))) = seq0((𝑐 ∈ 2𝑜, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚 ∈ (bits‘𝐴), 𝑚 ∈ (bits‘𝐵), ∅ ∈ 𝑐), 1𝑜, ∅)), (𝑛 ∈ ℕ0 ↦ if(𝑛 = 0, ∅, (𝑛 − 1))))
11 eqid 2622 . . . . . . . . 9 (bits ↾ ℕ0) = (bits ↾ ℕ0)
12 simpll 790 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝐴 ∈ ℤ)
13 simplr 792 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝐵 ∈ ℤ)
14 simpr 477 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ ℕ0)
15 1nn0 11308 . . . . . . . . . . 11 1 ∈ ℕ0
1615a1i 11 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 1 ∈ ℕ0)
1714, 16nn0addcld 11355 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝑘 + 1) ∈ ℕ0)
1810, 11, 12, 13, 17sadaddlem 15188 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (((bits‘𝐴) sadd (bits‘𝐵)) ∩ (0..^(𝑘 + 1))) = (bits‘((𝐴 + 𝐵) mod (2↑(𝑘 + 1)))))
1912, 13zaddcld 11486 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝐴 + 𝐵) ∈ ℤ)
20 bitsmod 15158 . . . . . . . . 9 (((𝐴 + 𝐵) ∈ ℤ ∧ (𝑘 + 1) ∈ ℕ0) → (bits‘((𝐴 + 𝐵) mod (2↑(𝑘 + 1)))) = ((bits‘(𝐴 + 𝐵)) ∩ (0..^(𝑘 + 1))))
2119, 17, 20syl2anc 693 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (bits‘((𝐴 + 𝐵) mod (2↑(𝑘 + 1)))) = ((bits‘(𝐴 + 𝐵)) ∩ (0..^(𝑘 + 1))))
2218, 21eqtrd 2656 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (((bits‘𝐴) sadd (bits‘𝐵)) ∩ (0..^(𝑘 + 1))) = ((bits‘(𝐴 + 𝐵)) ∩ (0..^(𝑘 + 1))))
2322eleq2d 2687 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝑘 ∈ (((bits‘𝐴) sadd (bits‘𝐵)) ∩ (0..^(𝑘 + 1))) ↔ 𝑘 ∈ ((bits‘(𝐴 + 𝐵)) ∩ (0..^(𝑘 + 1)))))
24 elin 3796 . . . . . 6 (𝑘 ∈ (((bits‘𝐴) sadd (bits‘𝐵)) ∩ (0..^(𝑘 + 1))) ↔ (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1))))
25 elin 3796 . . . . . 6 (𝑘 ∈ ((bits‘(𝐴 + 𝐵)) ∩ (0..^(𝑘 + 1))) ↔ (𝑘 ∈ (bits‘(𝐴 + 𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1))))
2623, 24, 253bitr3g 302 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → ((𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1))) ↔ (𝑘 ∈ (bits‘(𝐴 + 𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1)))))
27 nn0uz 11722 . . . . . . . . 9 0 = (ℤ‘0)
2814, 27syl6eleq 2711 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ (ℤ‘0))
29 eluzfz2 12349 . . . . . . . 8 (𝑘 ∈ (ℤ‘0) → 𝑘 ∈ (0...𝑘))
3028, 29syl 17 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ (0...𝑘))
3114nn0zd 11480 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ ℤ)
32 fzval3 12536 . . . . . . . 8 (𝑘 ∈ ℤ → (0...𝑘) = (0..^(𝑘 + 1)))
3331, 32syl 17 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (0...𝑘) = (0..^(𝑘 + 1)))
3430, 33eleqtrd 2703 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ (0..^(𝑘 + 1)))
3534biantrud 528 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ↔ (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1)))))
3634biantrud 528 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝑘 ∈ (bits‘(𝐴 + 𝐵)) ↔ (𝑘 ∈ (bits‘(𝐴 + 𝐵)) ∧ 𝑘 ∈ (0..^(𝑘 + 1)))))
3726, 35, 363bitr4d 300 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ 𝑘 ∈ ℕ0) → (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ↔ 𝑘 ∈ (bits‘(𝐴 + 𝐵))))
3837ex 450 . . 3 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑘 ∈ ℕ0 → (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ↔ 𝑘 ∈ (bits‘(𝐴 + 𝐵)))))
396, 9, 38pm5.21ndd 369 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑘 ∈ ((bits‘𝐴) sadd (bits‘𝐵)) ↔ 𝑘 ∈ (bits‘(𝐴 + 𝐵))))
4039eqrdv 2620 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((bits‘𝐴) sadd (bits‘𝐵)) = (bits‘(𝐴 + 𝐵)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 196  wa 384   = wceq 1483  caddwcad 1545  wcel 1990  cin 3573  wss 3574  c0 3915  ifcif 4086  cmpt 4729  ccnv 5113  cres 5116  cfv 5888  (class class class)co 6650  cmpt2 6652  1𝑜c1o 7553  2𝑜c2o 7554  0cc0 9936  1c1 9937   + caddc 9939  cmin 10266  2c2 11070  0cn0 11292  cz 11377  cuz 11687  ...cfz 12326  ..^cfzo 12465   mod cmo 12668  seqcseq 12801  cexp 12860  bitscbits 15141   sadd csad 15142
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-8 1992  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-rep 4771  ax-sep 4781  ax-nul 4789  ax-pow 4843  ax-pr 4906  ax-un 6949  ax-inf2 8538  ax-cnex 9992  ax-resscn 9993  ax-1cn 9994  ax-icn 9995  ax-addcl 9996  ax-addrcl 9997  ax-mulcl 9998  ax-mulrcl 9999  ax-mulcom 10000  ax-addass 10001  ax-mulass 10002  ax-distr 10003  ax-i2m1 10004  ax-1ne0 10005  ax-1rid 10006  ax-rnegex 10007  ax-rrecex 10008  ax-cnre 10009  ax-pre-lttri 10010  ax-pre-lttrn 10011  ax-pre-ltadd 10012  ax-pre-mulgt0 10013  ax-pre-sup 10014
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3or 1038  df-3an 1039  df-xor 1465  df-tru 1486  df-fal 1489  df-had 1533  df-cad 1546  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-ne 2795  df-nel 2898  df-ral 2917  df-rex 2918  df-reu 2919  df-rmo 2920  df-rab 2921  df-v 3202  df-sbc 3436  df-csb 3534  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-pss 3590  df-nul 3916  df-if 4087  df-pw 4160  df-sn 4178  df-pr 4180  df-tp 4182  df-op 4184  df-uni 4437  df-int 4476  df-iun 4522  df-disj 4621  df-br 4654  df-opab 4713  df-mpt 4730  df-tr 4753  df-id 5024  df-eprel 5029  df-po 5035  df-so 5036  df-fr 5073  df-se 5074  df-we 5075  df-xp 5120  df-rel 5121  df-cnv 5122  df-co 5123  df-dm 5124  df-rn 5125  df-res 5126  df-ima 5127  df-pred 5680  df-ord 5726  df-on 5727  df-lim 5728  df-suc 5729  df-iota 5851  df-fun 5890  df-fn 5891  df-f 5892  df-f1 5893  df-fo 5894  df-f1o 5895  df-fv 5896  df-isom 5897  df-riota 6611  df-ov 6653  df-oprab 6654  df-mpt2 6655  df-om 7066  df-1st 7168  df-2nd 7169  df-wrecs 7407  df-recs 7468  df-rdg 7506  df-1o 7560  df-2o 7561  df-oadd 7564  df-er 7742  df-map 7859  df-pm 7860  df-en 7956  df-dom 7957  df-sdom 7958  df-fin 7959  df-sup 8348  df-inf 8349  df-oi 8415  df-card 8765  df-cda 8990  df-pnf 10076  df-mnf 10077  df-xr 10078  df-ltxr 10079  df-le 10080  df-sub 10268  df-neg 10269  df-div 10685  df-nn 11021  df-2 11079  df-3 11080  df-n0 11293  df-xnn0 11364  df-z 11378  df-uz 11688  df-rp 11833  df-fz 12327  df-fzo 12466  df-fl 12593  df-mod 12669  df-seq 12802  df-exp 12861  df-hash 13118  df-cj 13839  df-re 13840  df-im 13841  df-sqrt 13975  df-abs 13976  df-clim 14219  df-sum 14417  df-dvds 14984  df-bits 15144  df-sad 15173
This theorem is referenced by:  bitsres  15195  smumullem  15214
  Copyright terms: Public domain W3C validator