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

Theorem saddisj 16447
Description: The sum of disjoint sequences is the union of the sequences. (In this case, there are no carried bits.) (Contributed by Mario Carneiro, 9-Sep-2016.)
Hypotheses
Ref Expression
saddisj.1 (𝜑𝐴 ⊆ ℕ0)
saddisj.2 (𝜑𝐵 ⊆ ℕ0)
saddisj.3 (𝜑 → (𝐴𝐵) = ∅)
Assertion
Ref Expression
saddisj (𝜑 → (𝐴 sadd 𝐵) = (𝐴𝐵))

Proof of Theorem saddisj
Dummy variables 𝑘 𝑐 𝑚 𝑥 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 saddisj.1 . . . . 5 (𝜑𝐴 ⊆ ℕ0)
2 saddisj.2 . . . . 5 (𝜑𝐵 ⊆ ℕ0)
3 sadcl 16444 . . . . 5 ((𝐴 ⊆ ℕ0𝐵 ⊆ ℕ0) → (𝐴 sadd 𝐵) ⊆ ℕ0)
41, 2, 3syl2anc 582 . . . 4 (𝜑 → (𝐴 sadd 𝐵) ⊆ ℕ0)
54sseld 3981 . . 3 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) → 𝑘 ∈ ℕ0))
61, 2unssd 4188 . . . 4 (𝜑 → (𝐴𝐵) ⊆ ℕ0)
76sseld 3981 . . 3 (𝜑 → (𝑘 ∈ (𝐴𝐵) → 𝑘 ∈ ℕ0))
81adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐴 ⊆ ℕ0)
92adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐵 ⊆ ℕ0)
10 saddisj.3 . . . . . 6 (𝜑 → (𝐴𝐵) = ∅)
1110adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → (𝐴𝐵) = ∅)
12 eqid 2728 . . . . 5 seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚𝐴, 𝑚𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1)))) = seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚𝐴, 𝑚𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1))))
13 simpr 483 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝑘 ∈ ℕ0)
148, 9, 11, 12, 13saddisjlem 16446 . . . 4 ((𝜑𝑘 ∈ ℕ0) → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1514ex 411 . . 3 (𝜑 → (𝑘 ∈ ℕ0 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵))))
165, 7, 15pm5.21ndd 378 . 2 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1716eqrdv 2726 1 (𝜑 → (𝐴 sadd 𝐵) = (𝐴𝐵))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 205  wa 394   = wceq 1533  caddwcad 1599  wcel 2098  cun 3947  cin 3948  wss 3949  c0 4326  ifcif 4532  cmpt 5235  (class class class)co 7426  cmpo 7428  1oc1o 8486  2oc2o 8487  0cc0 11146  1c1 11147  cmin 11482  0cn0 12510  seqcseq 14006   sadd csad 16402
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1789  ax-4 1803  ax-5 1905  ax-6 1963  ax-7 2003  ax-8 2100  ax-9 2108  ax-10 2129  ax-11 2146  ax-12 2166  ax-ext 2699  ax-sep 5303  ax-nul 5310  ax-pow 5369  ax-pr 5433  ax-un 7746  ax-cnex 11202  ax-resscn 11203  ax-1cn 11204  ax-icn 11205  ax-addcl 11206  ax-addrcl 11207  ax-mulcl 11208  ax-mulrcl 11209  ax-mulcom 11210  ax-addass 11211  ax-mulass 11212  ax-distr 11213  ax-i2m1 11214  ax-1ne0 11215  ax-1rid 11216  ax-rnegex 11217  ax-rrecex 11218  ax-cnre 11219  ax-pre-lttri 11220  ax-pre-lttrn 11221  ax-pre-ltadd 11222  ax-pre-mulgt0 11223
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 846  df-3or 1085  df-3an 1086  df-xor 1505  df-tru 1536  df-fal 1546  df-had 1587  df-cad 1600  df-ex 1774  df-nf 1778  df-sb 2060  df-mo 2529  df-eu 2558  df-clab 2706  df-cleq 2720  df-clel 2806  df-nfc 2881  df-ne 2938  df-nel 3044  df-ral 3059  df-rex 3068  df-reu 3375  df-rab 3431  df-v 3475  df-sbc 3779  df-csb 3895  df-dif 3952  df-un 3954  df-in 3956  df-ss 3966  df-pss 3968  df-nul 4327  df-if 4533  df-pw 4608  df-sn 4633  df-pr 4635  df-op 4639  df-uni 4913  df-iun 5002  df-br 5153  df-opab 5215  df-mpt 5236  df-tr 5270  df-id 5580  df-eprel 5586  df-po 5594  df-so 5595  df-fr 5637  df-we 5639  df-xp 5688  df-rel 5689  df-cnv 5690  df-co 5691  df-dm 5692  df-rn 5693  df-res 5694  df-ima 5695  df-pred 6310  df-ord 6377  df-on 6378  df-lim 6379  df-suc 6380  df-iota 6505  df-fun 6555  df-fn 6556  df-f 6557  df-f1 6558  df-fo 6559  df-f1o 6560  df-fv 6561  df-riota 7382  df-ov 7429  df-oprab 7430  df-mpo 7431  df-om 7877  df-1st 7999  df-2nd 8000  df-frecs 8293  df-wrecs 8324  df-recs 8398  df-rdg 8437  df-1o 8493  df-2o 8494  df-er 8731  df-en 8971  df-dom 8972  df-sdom 8973  df-pnf 11288  df-mnf 11289  df-xr 11290  df-ltxr 11291  df-le 11292  df-sub 11484  df-neg 11485  df-nn 12251  df-n0 12511  df-z 12597  df-uz 12861  df-fz 13525  df-seq 14007  df-sad 16433
This theorem is referenced by:  sadid1  16450  bitsres  16455
  Copyright terms: Public domain W3C validator