| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > fsumser | Structured version Visualization version GIF version | ||
| Description: A finite sum expressed in terms of a partial sum of an infinite series. The recursive definition follows as fsum1 15768 and fsump1i 15790, which should make our notation clear and from which, along with closure fsumcl 15754, we will derive the basic properties of finite sums. (Contributed by NM, 11-Dec-2005.) (Revised by Mario Carneiro, 21-Apr-2014.) |
| Ref | Expression |
|---|---|
| fsumser.1 | ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) = 𝐴) |
| fsumser.2 | ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) |
| fsumser.3 | ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → 𝐴 ∈ ℂ) |
| Ref | Expression |
|---|---|
| fsumser | ⊢ (𝜑 → Σ𝑘 ∈ (𝑀...𝑁)𝐴 = (seq𝑀( + , 𝐹)‘𝑁)) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | eleq1w 2818 | . . . . . 6 ⊢ (𝑚 = 𝑘 → (𝑚 ∈ (𝑀...𝑁) ↔ 𝑘 ∈ (𝑀...𝑁))) | |
| 2 | fveq2 6881 | . . . . . 6 ⊢ (𝑚 = 𝑘 → (𝐹‘𝑚) = (𝐹‘𝑘)) | |
| 3 | 1, 2 | ifbieq1d 4530 | . . . . 5 ⊢ (𝑚 = 𝑘 → if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0) = if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0)) |
| 4 | eqid 2736 | . . . . 5 ⊢ (𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0)) = (𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0)) | |
| 5 | fvex 6894 | . . . . . 6 ⊢ (𝐹‘𝑘) ∈ V | |
| 6 | c0ex 11234 | . . . . . 6 ⊢ 0 ∈ V | |
| 7 | 5, 6 | ifex 4556 | . . . . 5 ⊢ if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0) ∈ V |
| 8 | 3, 4, 7 | fvmpt 6991 | . . . 4 ⊢ (𝑘 ∈ (ℤ≥‘𝑀) → ((𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0))‘𝑘) = if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0)) |
| 9 | fsumser.1 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) = 𝐴) | |
| 10 | 9 | ifeq1da 4537 | . . . 4 ⊢ (𝜑 → if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0) = if(𝑘 ∈ (𝑀...𝑁), 𝐴, 0)) |
| 11 | 8, 10 | sylan9eqr 2793 | . . 3 ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → ((𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0))‘𝑘) = if(𝑘 ∈ (𝑀...𝑁), 𝐴, 0)) |
| 12 | fsumser.2 | . . 3 ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) | |
| 13 | fsumser.3 | . . 3 ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → 𝐴 ∈ ℂ) | |
| 14 | ssidd 3987 | . . 3 ⊢ (𝜑 → (𝑀...𝑁) ⊆ (𝑀...𝑁)) | |
| 15 | 11, 12, 13, 14 | fsumsers 15749 | . 2 ⊢ (𝜑 → Σ𝑘 ∈ (𝑀...𝑁)𝐴 = (seq𝑀( + , (𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0)))‘𝑁)) |
| 16 | elfzuz 13542 | . . . . . 6 ⊢ (𝑘 ∈ (𝑀...𝑁) → 𝑘 ∈ (ℤ≥‘𝑀)) | |
| 17 | 16, 8 | syl 17 | . . . . 5 ⊢ (𝑘 ∈ (𝑀...𝑁) → ((𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0))‘𝑘) = if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0)) |
| 18 | iftrue 4511 | . . . . 5 ⊢ (𝑘 ∈ (𝑀...𝑁) → if(𝑘 ∈ (𝑀...𝑁), (𝐹‘𝑘), 0) = (𝐹‘𝑘)) | |
| 19 | 17, 18 | eqtrd 2771 | . . . 4 ⊢ (𝑘 ∈ (𝑀...𝑁) → ((𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0))‘𝑘) = (𝐹‘𝑘)) |
| 20 | 19 | adantl 481 | . . 3 ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → ((𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0))‘𝑘) = (𝐹‘𝑘)) |
| 21 | 12, 20 | seqfveq 14049 | . 2 ⊢ (𝜑 → (seq𝑀( + , (𝑚 ∈ (ℤ≥‘𝑀) ↦ if(𝑚 ∈ (𝑀...𝑁), (𝐹‘𝑚), 0)))‘𝑁) = (seq𝑀( + , 𝐹)‘𝑁)) |
| 22 | 15, 21 | eqtrd 2771 | 1 ⊢ (𝜑 → Σ𝑘 ∈ (𝑀...𝑁)𝐴 = (seq𝑀( + , 𝐹)‘𝑁)) |
| Colors of variables: wff setvar class |
| Syntax hints: → wi 4 ∧ wa 395 = wceq 1540 ∈ wcel 2109 ifcif 4505 ↦ cmpt 5206 ‘cfv 6536 (class class class)co 7410 ℂcc 11132 0cc0 11134 + caddc 11137 ℤ≥cuz 12857 ...cfz 13529 seqcseq 14024 Σcsu 15707 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1795 ax-4 1809 ax-5 1910 ax-6 1967 ax-7 2008 ax-8 2111 ax-9 2119 ax-10 2142 ax-11 2158 ax-12 2178 ax-ext 2708 ax-rep 5254 ax-sep 5271 ax-nul 5281 ax-pow 5340 ax-pr 5407 ax-un 7734 ax-inf2 9660 ax-cnex 11190 ax-resscn 11191 ax-1cn 11192 ax-icn 11193 ax-addcl 11194 ax-addrcl 11195 ax-mulcl 11196 ax-mulrcl 11197 ax-mulcom 11198 ax-addass 11199 ax-mulass 11200 ax-distr 11201 ax-i2m1 11202 ax-1ne0 11203 ax-1rid 11204 ax-rnegex 11205 ax-rrecex 11206 ax-cnre 11207 ax-pre-lttri 11208 ax-pre-lttrn 11209 ax-pre-ltadd 11210 ax-pre-mulgt0 11211 ax-pre-sup 11212 |
| This theorem depends on definitions: df-bi 207 df-an 396 df-or 848 df-3or 1087 df-3an 1088 df-tru 1543 df-fal 1553 df-ex 1780 df-nf 1784 df-sb 2066 df-mo 2540 df-eu 2569 df-clab 2715 df-cleq 2728 df-clel 2810 df-nfc 2886 df-ne 2934 df-nel 3038 df-ral 3053 df-rex 3062 df-rmo 3364 df-reu 3365 df-rab 3421 df-v 3466 df-sbc 3771 df-csb 3880 df-dif 3934 df-un 3936 df-in 3938 df-ss 3948 df-pss 3951 df-nul 4314 df-if 4506 df-pw 4582 df-sn 4607 df-pr 4609 df-op 4613 df-uni 4889 df-int 4928 df-iun 4974 df-br 5125 df-opab 5187 df-mpt 5207 df-tr 5235 df-id 5553 df-eprel 5558 df-po 5566 df-so 5567 df-fr 5611 df-se 5612 df-we 5613 df-xp 5665 df-rel 5666 df-cnv 5667 df-co 5668 df-dm 5669 df-rn 5670 df-res 5671 df-ima 5672 df-pred 6295 df-ord 6360 df-on 6361 df-lim 6362 df-suc 6363 df-iota 6489 df-fun 6538 df-fn 6539 df-f 6540 df-f1 6541 df-fo 6542 df-f1o 6543 df-fv 6544 df-isom 6545 df-riota 7367 df-ov 7413 df-oprab 7414 df-mpo 7415 df-om 7867 df-1st 7993 df-2nd 7994 df-frecs 8285 df-wrecs 8316 df-recs 8390 df-rdg 8429 df-1o 8485 df-er 8724 df-en 8965 df-dom 8966 df-sdom 8967 df-fin 8968 df-sup 9459 df-oi 9529 df-card 9958 df-pnf 11276 df-mnf 11277 df-xr 11278 df-ltxr 11279 df-le 11280 df-sub 11473 df-neg 11474 df-div 11900 df-nn 12246 df-2 12308 df-3 12309 df-n0 12507 df-z 12594 df-uz 12858 df-rp 13014 df-fz 13530 df-fzo 13677 df-seq 14025 df-exp 14085 df-hash 14354 df-cj 15123 df-re 15124 df-im 15125 df-sqrt 15259 df-abs 15260 df-clim 15509 df-sum 15708 |
| This theorem is referenced by: isumclim3 15780 seqabs 15835 cvgcmpce 15839 isumsplit 15861 climcndslem1 15870 climcndslem2 15871 climcnds 15872 trireciplem 15883 geolim 15891 geo2lim 15896 mertenslem2 15906 mertens 15907 efcvgfsum 16107 effsumlt 16134 prmreclem6 16946 prmrec 16947 ovollb2lem 25446 ovoliunlem1 25460 ovoliun2 25464 ovolscalem1 25471 ovolicc2lem4 25478 uniioovol 25537 uniioombllem3 25543 uniioombllem6 25546 mtest 26370 mtestbdd 26371 psercn2 26389 psercn2OLD 26390 pserdvlem2 26395 abelthlem6 26403 logfac 26567 emcllem5 26967 lgamcvg2 27022 basellem8 27055 prmorcht 27145 pclogsum 27183 dchrisumlem2 27458 dchrmusum2 27462 dchrvmasumiflem1 27469 dchrisum0re 27481 dchrisum0lem1b 27483 dchrisum0lem2a 27485 dchrisum0lem2 27486 esumpcvgval 34114 esumcvg 34122 esumcvgsum 34124 knoppcnlem11 36526 fsumsermpt 45575 sumnnodd 45626 fourierdlem112 46214 sge0isum 46423 sge0seq 46442 |
| Copyright terms: Public domain | W3C validator |