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

Theorem saddisj 16443
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 16440 . . . . 5 ((𝐴 ⊆ ℕ0𝐵 ⊆ ℕ0) → (𝐴 sadd 𝐵) ⊆ ℕ0)
41, 2, 3syl2anc 582 . . . 4 (𝜑 → (𝐴 sadd 𝐵) ⊆ ℕ0)
54sseld 3975 . . 3 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) → 𝑘 ∈ ℕ0))
61, 2unssd 4184 . . . 4 (𝜑 → (𝐴𝐵) ⊆ ℕ0)
76sseld 3975 . . 3 (𝜑 → (𝑘 ∈ (𝐴𝐵) → 𝑘 ∈ ℕ0))
81adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐴 ⊆ ℕ0)
92adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐵 ⊆ ℕ0)
10 saddisj.3 . . . . . 6 (𝜑 → (𝐴𝐵) = ∅)
1110adantr 479 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → (𝐴𝐵) = ∅)
12 eqid 2725 . . . . 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 16442 . . . 4 ((𝜑𝑘 ∈ ℕ0) → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1514ex 411 . . 3 (𝜑 → (𝑘 ∈ ℕ0 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵))))
165, 7, 15pm5.21ndd 378 . 2 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1716eqrdv 2723 1 (𝜑 → (𝐴 sadd 𝐵) = (𝐴𝐵))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 205  wa 394   = wceq 1533  caddwcad 1599  wcel 2098  cun 3942  cin 3943  wss 3944  c0 4322  ifcif 4530  cmpt 5232  (class class class)co 7419  cmpo 7421  1oc1o 8480  2oc2o 8481  0cc0 11140  1c1 11141  cmin 11476  0cn0 12505  seqcseq 14002   sadd csad 16398
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 2696  ax-sep 5300  ax-nul 5307  ax-pow 5365  ax-pr 5429  ax-un 7741  ax-cnex 11196  ax-resscn 11197  ax-1cn 11198  ax-icn 11199  ax-addcl 11200  ax-addrcl 11201  ax-mulcl 11202  ax-mulrcl 11203  ax-mulcom 11204  ax-addass 11205  ax-mulass 11206  ax-distr 11207  ax-i2m1 11208  ax-1ne0 11209  ax-1rid 11210  ax-rnegex 11211  ax-rrecex 11212  ax-cnre 11213  ax-pre-lttri 11214  ax-pre-lttrn 11215  ax-pre-ltadd 11216  ax-pre-mulgt0 11217
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 2528  df-eu 2557  df-clab 2703  df-cleq 2717  df-clel 2802  df-nfc 2877  df-ne 2930  df-nel 3036  df-ral 3051  df-rex 3060  df-reu 3364  df-rab 3419  df-v 3463  df-sbc 3774  df-csb 3890  df-dif 3947  df-un 3949  df-in 3951  df-ss 3961  df-pss 3964  df-nul 4323  df-if 4531  df-pw 4606  df-sn 4631  df-pr 4633  df-op 4637  df-uni 4910  df-iun 4999  df-br 5150  df-opab 5212  df-mpt 5233  df-tr 5267  df-id 5576  df-eprel 5582  df-po 5590  df-so 5591  df-fr 5633  df-we 5635  df-xp 5684  df-rel 5685  df-cnv 5686  df-co 5687  df-dm 5688  df-rn 5689  df-res 5690  df-ima 5691  df-pred 6307  df-ord 6374  df-on 6375  df-lim 6376  df-suc 6377  df-iota 6501  df-fun 6551  df-fn 6552  df-f 6553  df-f1 6554  df-fo 6555  df-f1o 6556  df-fv 6557  df-riota 7375  df-ov 7422  df-oprab 7423  df-mpo 7424  df-om 7872  df-1st 7994  df-2nd 7995  df-frecs 8287  df-wrecs 8318  df-recs 8392  df-rdg 8431  df-1o 8487  df-2o 8488  df-er 8725  df-en 8965  df-dom 8966  df-sdom 8967  df-pnf 11282  df-mnf 11283  df-xr 11284  df-ltxr 11285  df-le 11286  df-sub 11478  df-neg 11479  df-nn 12246  df-n0 12506  df-z 12592  df-uz 12856  df-fz 13520  df-seq 14003  df-sad 16429
This theorem is referenced by:  sadid1  16446  bitsres  16451
  Copyright terms: Public domain W3C validator