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

Theorem saddisj 16410
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 16407 . . . . 5 ((𝐴 ⊆ ℕ0𝐵 ⊆ ℕ0) → (𝐴 sadd 𝐵) ⊆ ℕ0)
41, 2, 3syl2anc 583 . . . 4 (𝜑 → (𝐴 sadd 𝐵) ⊆ ℕ0)
54sseld 3976 . . 3 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) → 𝑘 ∈ ℕ0))
61, 2unssd 4181 . . . 4 (𝜑 → (𝐴𝐵) ⊆ ℕ0)
76sseld 3976 . . 3 (𝜑 → (𝑘 ∈ (𝐴𝐵) → 𝑘 ∈ ℕ0))
81adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐴 ⊆ ℕ0)
92adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐵 ⊆ ℕ0)
10 saddisj.3 . . . . . 6 (𝜑 → (𝐴𝐵) = ∅)
1110adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → (𝐴𝐵) = ∅)
12 eqid 2726 . . . . 5 seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚𝐴, 𝑚𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1)))) = seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚𝐴, 𝑚𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1))))
13 simpr 484 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝑘 ∈ ℕ0)
148, 9, 11, 12, 13saddisjlem 16409 . . . 4 ((𝜑𝑘 ∈ ℕ0) → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1514ex 412 . . 3 (𝜑 → (𝑘 ∈ ℕ0 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵))))
165, 7, 15pm5.21ndd 379 . 2 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1716eqrdv 2724 1 (𝜑 → (𝐴 sadd 𝐵) = (𝐴𝐵))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 205  wa 395   = wceq 1533  caddwcad 1599  wcel 2098  cun 3941  cin 3942  wss 3943  c0 4317  ifcif 4523  cmpt 5224  (class class class)co 7404  cmpo 7406  1oc1o 8457  2oc2o 8458  0cc0 11109  1c1 11110  cmin 11445  0cn0 12473  seqcseq 13969   sadd csad 16365
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 2163  ax-ext 2697  ax-sep 5292  ax-nul 5299  ax-pow 5356  ax-pr 5420  ax-un 7721  ax-cnex 11165  ax-resscn 11166  ax-1cn 11167  ax-icn 11168  ax-addcl 11169  ax-addrcl 11170  ax-mulcl 11171  ax-mulrcl 11172  ax-mulcom 11173  ax-addass 11174  ax-mulass 11175  ax-distr 11176  ax-i2m1 11177  ax-1ne0 11178  ax-1rid 11179  ax-rnegex 11180  ax-rrecex 11181  ax-cnre 11182  ax-pre-lttri 11183  ax-pre-lttrn 11184  ax-pre-ltadd 11185  ax-pre-mulgt0 11186
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 845  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 2528  df-eu 2557  df-clab 2704  df-cleq 2718  df-clel 2804  df-nfc 2879  df-ne 2935  df-nel 3041  df-ral 3056  df-rex 3065  df-reu 3371  df-rab 3427  df-v 3470  df-sbc 3773  df-csb 3889  df-dif 3946  df-un 3948  df-in 3950  df-ss 3960  df-pss 3962  df-nul 4318  df-if 4524  df-pw 4599  df-sn 4624  df-pr 4626  df-op 4630  df-uni 4903  df-iun 4992  df-br 5142  df-opab 5204  df-mpt 5225  df-tr 5259  df-id 5567  df-eprel 5573  df-po 5581  df-so 5582  df-fr 5624  df-we 5626  df-xp 5675  df-rel 5676  df-cnv 5677  df-co 5678  df-dm 5679  df-rn 5680  df-res 5681  df-ima 5682  df-pred 6293  df-ord 6360  df-on 6361  df-lim 6362  df-suc 6363  df-iota 6488  df-fun 6538  df-fn 6539  df-f 6540  df-f1 6541  df-fo 6542  df-f1o 6543  df-fv 6544  df-riota 7360  df-ov 7407  df-oprab 7408  df-mpo 7409  df-om 7852  df-1st 7971  df-2nd 7972  df-frecs 8264  df-wrecs 8295  df-recs 8369  df-rdg 8408  df-1o 8464  df-2o 8465  df-er 8702  df-en 8939  df-dom 8940  df-sdom 8941  df-pnf 11251  df-mnf 11252  df-xr 11253  df-ltxr 11254  df-le 11255  df-sub 11447  df-neg 11448  df-nn 12214  df-n0 12474  df-z 12560  df-uz 12824  df-fz 13488  df-seq 13970  df-sad 16396
This theorem is referenced by:  sadid1  16413  bitsres  16418
  Copyright terms: Public domain W3C validator