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

Theorem coemul 26207
Description: A coefficient of a product of polynomials. (Contributed by Mario Carneiro, 24-Jul-2014.)
Hypotheses
Ref Expression
coefv0.1 𝐴 = (coeff‘𝐹)
coeadd.2 𝐵 = (coeff‘𝐺)
Assertion
Ref Expression
coemul ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆) ∧ 𝑁 ∈ ℕ0) → ((coeff‘(𝐹f · 𝐺))‘𝑁) = Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
Distinct variable groups:   𝐴,𝑘   𝐵,𝑘   𝑘,𝐹   𝑘,𝐺   𝑘,𝑁   𝑆,𝑘

Proof of Theorem coemul
Dummy variable 𝑛 is distinct from all other variables.
StepHypRef Expression
1 coefv0.1 . . . . . 6 𝐴 = (coeff‘𝐹)
2 coeadd.2 . . . . . 6 𝐵 = (coeff‘𝐺)
3 eqid 2735 . . . . . 6 (deg‘𝐹) = (deg‘𝐹)
4 eqid 2735 . . . . . 6 (deg‘𝐺) = (deg‘𝐺)
51, 2, 3, 4coemullem 26205 . . . . 5 ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((coeff‘(𝐹f · 𝐺)) = (𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘)))) ∧ (deg‘(𝐹f · 𝐺)) ≤ ((deg‘𝐹) + (deg‘𝐺))))
65simpld 494 . . . 4 ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (coeff‘(𝐹f · 𝐺)) = (𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘)))))
76fveq1d 6877 . . 3 ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((coeff‘(𝐹f · 𝐺))‘𝑁) = ((𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘))))‘𝑁))
8 oveq2 7411 . . . . 5 (𝑛 = 𝑁 → (0...𝑛) = (0...𝑁))
9 fvoveq1 7426 . . . . . . 7 (𝑛 = 𝑁 → (𝐵‘(𝑛𝑘)) = (𝐵‘(𝑁𝑘)))
109oveq2d 7419 . . . . . 6 (𝑛 = 𝑁 → ((𝐴𝑘) · (𝐵‘(𝑛𝑘))) = ((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
1110adantr 480 . . . . 5 ((𝑛 = 𝑁𝑘 ∈ (0...𝑛)) → ((𝐴𝑘) · (𝐵‘(𝑛𝑘))) = ((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
128, 11sumeq12dv 15720 . . . 4 (𝑛 = 𝑁 → Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘))) = Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
13 eqid 2735 . . . 4 (𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘)))) = (𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘))))
14 sumex 15702 . . . 4 Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))) ∈ V
1512, 13, 14fvmpt 6985 . . 3 (𝑁 ∈ ℕ0 → ((𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘))))‘𝑁) = Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
167, 15sylan9eq 2790 . 2 (((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) ∧ 𝑁 ∈ ℕ0) → ((coeff‘(𝐹f · 𝐺))‘𝑁) = Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
17163impa 1109 1 ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆) ∧ 𝑁 ∈ ℕ0) → ((coeff‘(𝐹f · 𝐺))‘𝑁) = Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝐵‘(𝑁𝑘))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 395  w3a 1086   = wceq 1540  wcel 2108   class class class wbr 5119  cmpt 5201  cfv 6530  (class class class)co 7403  f cof 7667  0cc0 11127   + caddc 11130   · cmul 11132  cle 11268  cmin 11464  0cn0 12499  ...cfz 13522  Σcsu 15700  Polycply 26139  coeffccoe 26141  degcdgr 26142
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 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2157  ax-12 2177  ax-ext 2707  ax-rep 5249  ax-sep 5266  ax-nul 5276  ax-pow 5335  ax-pr 5402  ax-un 7727  ax-inf2 9653  ax-cnex 11183  ax-resscn 11184  ax-1cn 11185  ax-icn 11186  ax-addcl 11187  ax-addrcl 11188  ax-mulcl 11189  ax-mulrcl 11190  ax-mulcom 11191  ax-addass 11192  ax-mulass 11193  ax-distr 11194  ax-i2m1 11195  ax-1ne0 11196  ax-1rid 11197  ax-rnegex 11198  ax-rrecex 11199  ax-cnre 11200  ax-pre-lttri 11201  ax-pre-lttrn 11202  ax-pre-ltadd 11203  ax-pre-mulgt0 11204  ax-pre-sup 11205
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 2065  df-mo 2539  df-eu 2568  df-clab 2714  df-cleq 2727  df-clel 2809  df-nfc 2885  df-ne 2933  df-nel 3037  df-ral 3052  df-rex 3061  df-rmo 3359  df-reu 3360  df-rab 3416  df-v 3461  df-sbc 3766  df-csb 3875  df-dif 3929  df-un 3931  df-in 3933  df-ss 3943  df-pss 3946  df-nul 4309  df-if 4501  df-pw 4577  df-sn 4602  df-pr 4604  df-op 4608  df-uni 4884  df-int 4923  df-iun 4969  df-br 5120  df-opab 5182  df-mpt 5202  df-tr 5230  df-id 5548  df-eprel 5553  df-po 5561  df-so 5562  df-fr 5606  df-se 5607  df-we 5608  df-xp 5660  df-rel 5661  df-cnv 5662  df-co 5663  df-dm 5664  df-rn 5665  df-res 5666  df-ima 5667  df-pred 6290  df-ord 6355  df-on 6356  df-lim 6357  df-suc 6358  df-iota 6483  df-fun 6532  df-fn 6533  df-f 6534  df-f1 6535  df-fo 6536  df-f1o 6537  df-fv 6538  df-isom 6539  df-riota 7360  df-ov 7406  df-oprab 7407  df-mpo 7408  df-of 7669  df-om 7860  df-1st 7986  df-2nd 7987  df-frecs 8278  df-wrecs 8309  df-recs 8383  df-rdg 8422  df-1o 8478  df-er 8717  df-map 8840  df-pm 8841  df-en 8958  df-dom 8959  df-sdom 8960  df-fin 8961  df-sup 9452  df-inf 9453  df-oi 9522  df-card 9951  df-pnf 11269  df-mnf 11270  df-xr 11271  df-ltxr 11272  df-le 11273  df-sub 11466  df-neg 11467  df-div 11893  df-nn 12239  df-2 12301  df-3 12302  df-n0 12500  df-z 12587  df-uz 12851  df-rp 13007  df-fz 13523  df-fzo 13670  df-fl 13807  df-seq 14018  df-exp 14078  df-hash 14347  df-cj 15116  df-re 15117  df-im 15118  df-sqrt 15252  df-abs 15253  df-clim 15502  df-rlim 15503  df-sum 15701  df-0p 25621  df-ply 26143  df-coe 26145  df-dgr 26146
This theorem is referenced by:  coemulhi  26209  coemulc  26210  vieta1lem2  26269  plymulx0  34525
  Copyright terms: Public domain W3C validator