![]() |
Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > MPE Home > Th. List > saddisj | Structured version Visualization version GIF version |
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.) |
Ref | Expression |
---|---|
saddisj.1 | ⊢ (𝜑 → 𝐴 ⊆ ℕ0) |
saddisj.2 | ⊢ (𝜑 → 𝐵 ⊆ ℕ0) |
saddisj.3 | ⊢ (𝜑 → (𝐴 ∩ 𝐵) = ∅) |
Ref | Expression |
---|---|
saddisj | ⊢ (𝜑 → (𝐴 sadd 𝐵) = (𝐴 ∪ 𝐵)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | saddisj.1 | . . . . 5 ⊢ (𝜑 → 𝐴 ⊆ ℕ0) | |
2 | saddisj.2 | . . . . 5 ⊢ (𝜑 → 𝐵 ⊆ ℕ0) | |
3 | sadcl 15590 | . . . . 5 ⊢ ((𝐴 ⊆ ℕ0 ∧ 𝐵 ⊆ ℕ0) → (𝐴 sadd 𝐵) ⊆ ℕ0) | |
4 | 1, 2, 3 | syl2anc 579 | . . . 4 ⊢ (𝜑 → (𝐴 sadd 𝐵) ⊆ ℕ0) |
5 | 4 | sseld 3820 | . . 3 ⊢ (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) → 𝑘 ∈ ℕ0)) |
6 | 1, 2 | unssd 4012 | . . . 4 ⊢ (𝜑 → (𝐴 ∪ 𝐵) ⊆ ℕ0) |
7 | 6 | sseld 3820 | . . 3 ⊢ (𝜑 → (𝑘 ∈ (𝐴 ∪ 𝐵) → 𝑘 ∈ ℕ0)) |
8 | 1 | adantr 474 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ0) → 𝐴 ⊆ ℕ0) |
9 | 2 | adantr 474 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ0) → 𝐵 ⊆ ℕ0) |
10 | saddisj.3 | . . . . . 6 ⊢ (𝜑 → (𝐴 ∩ 𝐵) = ∅) | |
11 | 10 | adantr 474 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ0) → (𝐴 ∩ 𝐵) = ∅) |
12 | eqid 2778 | . . . . 5 ⊢ seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚 ∈ 𝐴, 𝑚 ∈ 𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1)))) = seq0((𝑐 ∈ 2o, 𝑚 ∈ ℕ0 ↦ if(cadd(𝑚 ∈ 𝐴, 𝑚 ∈ 𝐵, ∅ ∈ 𝑐), 1o, ∅)), (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (𝑥 − 1)))) | |
13 | simpr 479 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ0) → 𝑘 ∈ ℕ0) | |
14 | 8, 9, 11, 12, 13 | saddisjlem 15592 | . . . 4 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ0) → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴 ∪ 𝐵))) |
15 | 14 | ex 403 | . . 3 ⊢ (𝜑 → (𝑘 ∈ ℕ0 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴 ∪ 𝐵)))) |
16 | 5, 7, 15 | pm5.21ndd 371 | . 2 ⊢ (𝜑 → (𝑘 ∈ (𝐴 sadd 𝐵) ↔ 𝑘 ∈ (𝐴 ∪ 𝐵))) |
17 | 16 | eqrdv 2776 | 1 ⊢ (𝜑 → (𝐴 sadd 𝐵) = (𝐴 ∪ 𝐵)) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ↔ wb 198 ∧ wa 386 = wceq 1601 caddwcad 1664 ∈ wcel 2107 ∪ cun 3790 ∩ cin 3791 ⊆ wss 3792 ∅c0 4141 ifcif 4307 ↦ cmpt 4965 (class class class)co 6922 ↦ cmpt2 6924 1oc1o 7836 2oc2o 7837 0cc0 10272 1c1 10273 − cmin 10606 ℕ0cn0 11642 seqcseq 13119 sadd csad 15548 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1839 ax-4 1853 ax-5 1953 ax-6 2021 ax-7 2055 ax-8 2109 ax-9 2116 ax-10 2135 ax-11 2150 ax-12 2163 ax-13 2334 ax-ext 2754 ax-sep 5017 ax-nul 5025 ax-pow 5077 ax-pr 5138 ax-un 7226 ax-cnex 10328 ax-resscn 10329 ax-1cn 10330 ax-icn 10331 ax-addcl 10332 ax-addrcl 10333 ax-mulcl 10334 ax-mulrcl 10335 ax-mulcom 10336 ax-addass 10337 ax-mulass 10338 ax-distr 10339 ax-i2m1 10340 ax-1ne0 10341 ax-1rid 10342 ax-rnegex 10343 ax-rrecex 10344 ax-cnre 10345 ax-pre-lttri 10346 ax-pre-lttrn 10347 ax-pre-ltadd 10348 ax-pre-mulgt0 10349 |
This theorem depends on definitions: df-bi 199 df-an 387 df-or 837 df-3or 1072 df-3an 1073 df-xor 1583 df-tru 1605 df-had 1652 df-cad 1665 df-ex 1824 df-nf 1828 df-sb 2012 df-mo 2551 df-eu 2587 df-clab 2764 df-cleq 2770 df-clel 2774 df-nfc 2921 df-ne 2970 df-nel 3076 df-ral 3095 df-rex 3096 df-reu 3097 df-rab 3099 df-v 3400 df-sbc 3653 df-csb 3752 df-dif 3795 df-un 3797 df-in 3799 df-ss 3806 df-pss 3808 df-nul 4142 df-if 4308 df-pw 4381 df-sn 4399 df-pr 4401 df-tp 4403 df-op 4405 df-uni 4672 df-iun 4755 df-br 4887 df-opab 4949 df-mpt 4966 df-tr 4988 df-id 5261 df-eprel 5266 df-po 5274 df-so 5275 df-fr 5314 df-we 5316 df-xp 5361 df-rel 5362 df-cnv 5363 df-co 5364 df-dm 5365 df-rn 5366 df-res 5367 df-ima 5368 df-pred 5933 df-ord 5979 df-on 5980 df-lim 5981 df-suc 5982 df-iota 6099 df-fun 6137 df-fn 6138 df-f 6139 df-f1 6140 df-fo 6141 df-f1o 6142 df-fv 6143 df-riota 6883 df-ov 6925 df-oprab 6926 df-mpt2 6927 df-om 7344 df-1st 7445 df-2nd 7446 df-wrecs 7689 df-recs 7751 df-rdg 7789 df-1o 7843 df-2o 7844 df-er 8026 df-en 8242 df-dom 8243 df-sdom 8244 df-pnf 10413 df-mnf 10414 df-xr 10415 df-ltxr 10416 df-le 10417 df-sub 10608 df-neg 10609 df-nn 11375 df-n0 11643 df-z 11729 df-uz 11993 df-fz 12644 df-seq 13120 df-sad 15579 |
This theorem is referenced by: sadid1 15596 bitsres 15601 |
Copyright terms: Public domain | W3C validator |