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

Theorem saddisj 16499
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 16496 . . . . 5 ((𝐴 ⊆ ℕ0𝐵 ⊆ ℕ0) → (𝐴 sadd 𝐵) ⊆ ℕ0)
41, 2, 3syl2anc 584 . . . 4 (𝜑 → (𝐴 sadd 𝐵) ⊆ ℕ0)
54sseld 3994 . . 3 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) → 𝑘 ∈ ℕ0))
61, 2unssd 4202 . . . 4 (𝜑 → (𝐴𝐵) ⊆ ℕ0)
76sseld 3994 . . 3 (𝜑 → (𝑘 ∈ (𝐴𝐵) → 𝑘 ∈ ℕ0))
81adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐴 ⊆ ℕ0)
92adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → 𝐵 ⊆ ℕ0)
10 saddisj.3 . . . . . 6 (𝜑 → (𝐴𝐵) = ∅)
1110adantr 480 . . . . 5 ((𝜑𝑘 ∈ ℕ0) → (𝐴𝐵) = ∅)
12 eqid 2735 . . . . 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 16498 . . . 4 ((𝜑𝑘 ∈ ℕ0) → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1514ex 412 . . 3 (𝜑 → (𝑘 ∈ ℕ0 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵))))
165, 7, 15pm5.21ndd 379 . 2 (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴𝐵)))
1716eqrdv 2733 1 (𝜑 → (𝐴 sadd 𝐵) = (𝐴𝐵))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 206  wa 395   = wceq 1537  caddwcad 1603  wcel 2106  cun 3961  cin 3962  wss 3963  c0 4339  ifcif 4531  cmpt 5231  (class class class)co 7431  cmpo 7433  1oc1o 8498  2oc2o 8499  0cc0 11153  1c1 11154  cmin 11490  0cn0 12524  seqcseq 14039   sadd csad 16454
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1792  ax-4 1806  ax-5 1908  ax-6 1965  ax-7 2005  ax-8 2108  ax-9 2116  ax-10 2139  ax-11 2155  ax-12 2175  ax-ext 2706  ax-sep 5302  ax-nul 5312  ax-pow 5371  ax-pr 5438  ax-un 7754  ax-cnex 11209  ax-resscn 11210  ax-1cn 11211  ax-icn 11212  ax-addcl 11213  ax-addrcl 11214  ax-mulcl 11215  ax-mulrcl 11216  ax-mulcom 11217  ax-addass 11218  ax-mulass 11219  ax-distr 11220  ax-i2m1 11221  ax-1ne0 11222  ax-1rid 11223  ax-rnegex 11224  ax-rrecex 11225  ax-cnre 11226  ax-pre-lttri 11227  ax-pre-lttrn 11228  ax-pre-ltadd 11229  ax-pre-mulgt0 11230
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 848  df-3or 1087  df-3an 1088  df-xor 1509  df-tru 1540  df-fal 1550  df-had 1591  df-cad 1604  df-ex 1777  df-nf 1781  df-sb 2063  df-mo 2538  df-eu 2567  df-clab 2713  df-cleq 2727  df-clel 2814  df-nfc 2890  df-ne 2939  df-nel 3045  df-ral 3060  df-rex 3069  df-reu 3379  df-rab 3434  df-v 3480  df-sbc 3792  df-csb 3909  df-dif 3966  df-un 3968  df-in 3970  df-ss 3980  df-pss 3983  df-nul 4340  df-if 4532  df-pw 4607  df-sn 4632  df-pr 4634  df-op 4638  df-uni 4913  df-iun 4998  df-br 5149  df-opab 5211  df-mpt 5232  df-tr 5266  df-id 5583  df-eprel 5589  df-po 5597  df-so 5598  df-fr 5641  df-we 5643  df-xp 5695  df-rel 5696  df-cnv 5697  df-co 5698  df-dm 5699  df-rn 5700  df-res 5701  df-ima 5702  df-pred 6323  df-ord 6389  df-on 6390  df-lim 6391  df-suc 6392  df-iota 6516  df-fun 6565  df-fn 6566  df-f 6567  df-f1 6568  df-fo 6569  df-f1o 6570  df-fv 6571  df-riota 7388  df-ov 7434  df-oprab 7435  df-mpo 7436  df-om 7888  df-1st 8013  df-2nd 8014  df-frecs 8305  df-wrecs 8336  df-recs 8410  df-rdg 8449  df-1o 8505  df-2o 8506  df-er 8744  df-en 8985  df-dom 8986  df-sdom 8987  df-pnf 11295  df-mnf 11296  df-xr 11297  df-ltxr 11298  df-le 11299  df-sub 11492  df-neg 11493  df-nn 12265  df-n0 12525  df-z 12612  df-uz 12877  df-fz 13545  df-seq 14040  df-sad 16485
This theorem is referenced by:  sadid1  16502  bitsres  16507
  Copyright terms: Public domain W3C validator