![]() |
Metamath
Proof Explorer Theorem List (p. 241 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | plyssc 24001 | Every polynomial ring is contained in the ring of polynomials over ℂ. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (Poly‘𝑆) ⊆ (Poly‘ℂ) | ||
Theorem | elplyr 24002* | Sufficient condition for elementhood in the set of polynomials. (Contributed by Mario Carneiro, 17-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ ((𝑆 ⊆ ℂ ∧ 𝑁 ∈ ℕ0 ∧ 𝐴:ℕ0⟶𝑆) → (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘))) ∈ (Poly‘𝑆)) | ||
Theorem | elplyd 24003* | Sufficient condition for elementhood in the set of polynomials. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ ((𝜑 ∧ 𝑘 ∈ (0...𝑁)) → 𝐴 ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(𝐴 · (𝑧↑𝑘))) ∈ (Poly‘𝑆)) | ||
Theorem | ply1termlem 24004* | Lemma for ply1term 24005. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ 𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧↑𝑁))) ⇒ ⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(if(𝑘 = 𝑁, 𝐴, 0) · (𝑧↑𝑘)))) | ||
Theorem | ply1term 24005* | A one-term polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ 𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧↑𝑁))) ⇒ ⊢ ((𝑆 ⊆ ℂ ∧ 𝐴 ∈ 𝑆 ∧ 𝑁 ∈ ℕ0) → 𝐹 ∈ (Poly‘𝑆)) | ||
Theorem | plypow 24006* | A power is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ ((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆 ∧ 𝑁 ∈ ℕ0) → (𝑧 ∈ ℂ ↦ (𝑧↑𝑁)) ∈ (Poly‘𝑆)) | ||
Theorem | plyconst 24007 | A constant function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ ((𝑆 ⊆ ℂ ∧ 𝐴 ∈ 𝑆) → (ℂ × {𝐴}) ∈ (Poly‘𝑆)) | ||
Theorem | ne0p 24008 | A test to show that a polynomial is nonzero. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ ((𝐴 ∈ ℂ ∧ (𝐹‘𝐴) ≠ 0) → 𝐹 ≠ 0𝑝) | ||
Theorem | ply0 24009 | The zero function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ (𝑆 ⊆ ℂ → 0𝑝 ∈ (Poly‘𝑆)) | ||
Theorem | plyid 24010 | The identity function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.) |
⊢ ((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆) → Xp ∈ (Poly‘𝑆)) | ||
Theorem | plyeq0lem 24011* | Lemma for plyeq0 24012. If 𝐴 is the coefficient function for a nonzero polynomial such that 𝑃(𝑧) = Σ𝑘 ∈ ℕ0𝐴(𝑘) · 𝑧↑𝑘 = 0 for every 𝑧 ∈ ℂ and 𝐴(𝑀) is the nonzero leading coefficient, then the function 𝐹(𝑧) = 𝑃(𝑧) / 𝑧↑𝑀 is a sum of powers of 1 / 𝑧, and so the limit of this function as 𝑧 ⇝ +∞ is the constant term, 𝐴(𝑀). But 𝐹(𝑧) = 0 everywhere, so this limit is also equal to zero so that 𝐴(𝑀) = 0, a contradiction. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 0𝑝 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ 𝑀 = sup((◡𝐴 “ (𝑆 ∖ {0})), ℝ, < ) & ⊢ (𝜑 → (◡𝐴 “ (𝑆 ∖ {0})) ≠ ∅) ⇒ ⊢ ¬ 𝜑 | ||
Theorem | plyeq0 24012* | If a polynomial is zero at every point (or even just zero at the positive integers), then all the coefficients must be zero. This is the basis for the method of equating coefficients of equal polynomials, and ensures that df-coe 23991 is well-defined. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 0𝑝 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → 𝐴 = (ℕ0 × {0})) | ||
Theorem | plypf1 24013 | Write the set of complex polynomials in a subring in terms of the abstract polynomial construction. (Contributed by Mario Carneiro, 3-Jul-2015.) (Proof shortened by AV, 29-Sep-2019.) |
⊢ 𝑅 = (ℂfld ↾s 𝑆) & ⊢ 𝑃 = (Poly1‘𝑅) & ⊢ 𝐴 = (Base‘𝑃) & ⊢ 𝐸 = (eval1‘ℂfld) ⇒ ⊢ (𝑆 ∈ (SubRing‘ℂfld) → (Poly‘𝑆) = (𝐸 “ 𝐴)) | ||
Theorem | plyaddlem1 24014* | Derive the coefficient function for the sum of two polynomials. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴:ℕ0⟶ℂ) & ⊢ (𝜑 → 𝐵:ℕ0⟶ℂ) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 + 𝐺) = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...if(𝑀 ≤ 𝑁, 𝑁, 𝑀))(((𝐴 ∘𝑓 + 𝐵)‘𝑘) · (𝑧↑𝑘)))) | ||
Theorem | plymullem1 24015* | Derive the coefficient function for the product of two polynomials. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴:ℕ0⟶ℂ) & ⊢ (𝜑 → 𝐵:ℕ0⟶ℂ) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 · 𝐺) = (𝑧 ∈ ℂ ↦ Σ𝑛 ∈ (0...(𝑀 + 𝑁))(Σ𝑘 ∈ (0...𝑛)((𝐴‘𝑘) · (𝐵‘(𝑛 − 𝑘))) · (𝑧↑𝑛)))) | ||
Theorem | plyaddlem 24016* | Lemma for plyadd 24018. (Contributed by Mario Carneiro, 21-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → 𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 + 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | plymullem 24017* | Lemma for plymul 24019. (Contributed by Mario Carneiro, 21-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → 𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵‘𝑘) · (𝑧↑𝑘)))) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 · 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | plyadd 24018* | The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 + 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | plymul 24019* | The product of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 · 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | plysub 24020* | The difference of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 − 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | plyaddcl 24021 | The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹 ∘𝑓 + 𝐺) ∈ (Poly‘ℂ)) | ||
Theorem | plymulcl 24022 | The product of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹 ∘𝑓 · 𝐺) ∈ (Poly‘ℂ)) | ||
Theorem | plysubcl 24023 | The difference of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹 ∘𝑓 − 𝐺) ∈ (Poly‘ℂ)) | ||
Theorem | coeval 24024* | Value of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → (coeff‘𝐹) = (℩𝑎 ∈ (ℂ ↑𝑚 ℕ0)∃𝑛 ∈ ℕ0 ((𝑎 “ (ℤ≥‘(𝑛 + 1))) = {0} ∧ 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎‘𝑘) · (𝑧↑𝑘)))))) | ||
Theorem | coeeulem 24025* | Lemma for coeeu 24026. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐴 ∈ (ℂ ↑𝑚 ℕ0)) & ⊢ (𝜑 → 𝐵 ∈ (ℂ ↑𝑚 ℕ0)) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) | ||
Theorem | coeeu 24026* | Uniqueness of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → ∃!𝑎 ∈ (ℂ ↑𝑚 ℕ0)∃𝑛 ∈ ℕ0 ((𝑎 “ (ℤ≥‘(𝑛 + 1))) = {0} ∧ 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎‘𝑘) · (𝑧↑𝑘))))) | ||
Theorem | coelem 24027* | Lemma for properties of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → ((coeff‘𝐹) ∈ (ℂ ↑𝑚 ℕ0) ∧ ∃𝑛 ∈ ℕ0 (((coeff‘𝐹) “ (ℤ≥‘(𝑛 + 1))) = {0} ∧ 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)(((coeff‘𝐹)‘𝑘) · (𝑧↑𝑘)))))) | ||
Theorem | coeeq 24028* | If 𝐴 satisfies the properties of the coefficient function, it must be equal to the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴:ℕ0⟶ℂ) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (coeff‘𝐹) = 𝐴) | ||
Theorem | dgrval 24029 | Value of the degree function. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (deg‘𝐹) = sup((◡𝐴 “ (ℂ ∖ {0})), ℕ0, < )) | ||
Theorem | dgrlem 24030* | Lemma for dgrcl 24034 and similar theorems. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (𝐴:ℕ0⟶(𝑆 ∪ {0}) ∧ ∃𝑛 ∈ ℤ ∀𝑥 ∈ (◡𝐴 “ (ℂ ∖ {0}))𝑥 ≤ 𝑛)) | ||
Theorem | coef 24031 | The domain and range of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐴:ℕ0⟶(𝑆 ∪ {0})) | ||
Theorem | coef2 24032 | The domain and range of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 0 ∈ 𝑆) → 𝐴:ℕ0⟶𝑆) | ||
Theorem | coef3 24033 | The domain and range of the coefficient function. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐴:ℕ0⟶ℂ) | ||
Theorem | dgrcl 24034 | The degree of any polynomial is a nonnegative integer. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → (deg‘𝐹) ∈ ℕ0) | ||
Theorem | dgrub 24035 | If the 𝑀-th coefficient of 𝐹 is nonzero, then the degree of 𝐹 is at least 𝑀. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑀 ∈ ℕ0 ∧ (𝐴‘𝑀) ≠ 0) → 𝑀 ≤ 𝑁) | ||
Theorem | dgrub2 24036 | All the coefficients above the degree of 𝐹 are zero. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (𝐴 “ (ℤ≥‘(𝑁 + 1))) = {0}) | ||
Theorem | dgrlb 24037 | If all the coefficients above 𝑀 are zero, then the degree of 𝐹 is at most 𝑀. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑀 ∈ ℕ0 ∧ (𝐴 “ (ℤ≥‘(𝑀 + 1))) = {0}) → 𝑁 ≤ 𝑀) | ||
Theorem | coeidlem 24038* | Lemma for coeid 24039. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚 ℕ0)) & ⊢ (𝜑 → (𝐵 “ (ℤ≥‘(𝑀 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐵‘𝑘) · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) | ||
Theorem | coeid 24039* | Reconstruct a polynomial as an explicit sum of the coefficient function up to the degree of the polynomial. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) | ||
Theorem | coeid2 24040* | Reconstruct a polynomial as an explicit sum of the coefficient function up to the degree of the polynomial. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑋 ∈ ℂ) → (𝐹‘𝑋) = Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑋↑𝑘))) | ||
Theorem | coeid3 24041* | Reconstruct a polynomial as an explicit sum of the coefficient function up to at least the degree of the polynomial. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝑁 = (deg‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑀 ∈ (ℤ≥‘𝑁) ∧ 𝑋 ∈ ℂ) → (𝐹‘𝑋) = Σ𝑘 ∈ (0...𝑀)((𝐴‘𝑘) · (𝑋↑𝑘))) | ||
Theorem | plyco 24042* | The composition of two polynomials is a polynomial. (Contributed by Mario Carneiro, 23-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐹 ∘ 𝐺) ∈ (Poly‘𝑆)) | ||
Theorem | coeeq2 24043* | Compute the coefficient function given a sum expression for the polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ ((𝜑 ∧ 𝑘 ∈ (0...𝑁)) → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(𝐴 · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (coeff‘𝐹) = (𝑘 ∈ ℕ0 ↦ if(𝑘 ≤ 𝑁, 𝐴, 0))) | ||
Theorem | dgrle 24044* | Given an explicit expression for a polynomial, the degree is at most the highest term in the sum. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ ((𝜑 ∧ 𝑘 ∈ (0...𝑁)) → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(𝐴 · (𝑧↑𝑘)))) ⇒ ⊢ (𝜑 → (deg‘𝐹) ≤ 𝑁) | ||
Theorem | dgreq 24045* | If the highest term in a polynomial expression is nonzero, then the polynomial's degree is completely determined. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐴:ℕ0⟶ℂ) & ⊢ (𝜑 → (𝐴 “ (ℤ≥‘(𝑁 + 1))) = {0}) & ⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → (𝐴‘𝑁) ≠ 0) ⇒ ⊢ (𝜑 → (deg‘𝐹) = 𝑁) | ||
Theorem | 0dgr 24046 | A constant function has degree 0. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ (𝐴 ∈ ℂ → (deg‘(ℂ × {𝐴})) = 0) | ||
Theorem | 0dgrb 24047 | A function has degree zero iff it is a constant function. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → ((deg‘𝐹) = 0 ↔ 𝐹 = (ℂ × {(𝐹‘0)}))) | ||
Theorem | dgrnznn 24048 | A nonzero polynomial with a root has positive degree. (Contributed by Stefan O'Rear, 25-Nov-2014.) |
⊢ (((𝑃 ∈ (Poly‘𝑆) ∧ 𝑃 ≠ 0𝑝) ∧ (𝐴 ∈ ℂ ∧ (𝑃‘𝐴) = 0)) → (deg‘𝑃) ∈ ℕ) | ||
Theorem | coefv0 24049 | The result of evaluating a polynomial at zero is the constant term. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (𝐹‘0) = (𝐴‘0)) | ||
Theorem | coeaddlem 24050 | Lemma for coeadd 24052 and dgradd 24068. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) & ⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((coeff‘(𝐹 ∘𝑓 + 𝐺)) = (𝐴 ∘𝑓 + 𝐵) ∧ (deg‘(𝐹 ∘𝑓 + 𝐺)) ≤ if(𝑀 ≤ 𝑁, 𝑁, 𝑀))) | ||
Theorem | coemullem 24051* | Lemma for coemul 24053 and dgrmul 24071. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) & ⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((coeff‘(𝐹 ∘𝑓 · 𝐺)) = (𝑛 ∈ ℕ0 ↦ Σ𝑘 ∈ (0...𝑛)((𝐴‘𝑘) · (𝐵‘(𝑛 − 𝑘)))) ∧ (deg‘(𝐹 ∘𝑓 · 𝐺)) ≤ (𝑀 + 𝑁))) | ||
Theorem | coeadd 24052 | The coefficient function of a sum is the sum of coefficients. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (coeff‘(𝐹 ∘𝑓 + 𝐺)) = (𝐴 ∘𝑓 + 𝐵)) | ||
Theorem | coemul 24053* | A coefficient of a product of polynomials. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆) ∧ 𝑁 ∈ ℕ0) → ((coeff‘(𝐹 ∘𝑓 · 𝐺))‘𝑁) = Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝐵‘(𝑁 − 𝑘)))) | ||
Theorem | coe11 24054 | The coefficient function is one-to-one, so if the coefficients are equal then the functions are equal and vice-versa. (Contributed by Mario Carneiro, 24-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹 = 𝐺 ↔ 𝐴 = 𝐵)) | ||
Theorem | coemulhi 24055 | The leading coefficient of a product of polynomials. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) & ⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((coeff‘(𝐹 ∘𝑓 · 𝐺))‘(𝑀 + 𝑁)) = ((𝐴‘𝑀) · (𝐵‘𝑁))) | ||
Theorem | coemulc 24056 | The coefficient function is linear under scalar multiplication. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐹 ∈ (Poly‘𝑆)) → (coeff‘((ℂ × {𝐴}) ∘𝑓 · 𝐹)) = ((ℕ0 × {𝐴}) ∘𝑓 · (coeff‘𝐹))) | ||
Theorem | coe0 24057 | The coefficients of the zero polynomial are zero. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (coeff‘0𝑝) = (ℕ0 × {0}) | ||
Theorem | coesub 24058 | The coefficient function of a sum is the sum of coefficients. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (coeff‘(𝐹 ∘𝑓 − 𝐺)) = (𝐴 ∘𝑓 − 𝐵)) | ||
Theorem | coe1termlem 24059* | The coefficient function of a monomial. (Contributed by Mario Carneiro, 26-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) |
⊢ 𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧↑𝑁))) ⇒ ⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → ((coeff‘𝐹) = (𝑛 ∈ ℕ0 ↦ if(𝑛 = 𝑁, 𝐴, 0)) ∧ (𝐴 ≠ 0 → (deg‘𝐹) = 𝑁))) | ||
Theorem | coe1term 24060* | The coefficient function of a monomial. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ 𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧↑𝑁))) ⇒ ⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0 ∧ 𝑀 ∈ ℕ0) → ((coeff‘𝐹)‘𝑀) = if(𝑀 = 𝑁, 𝐴, 0)) | ||
Theorem | dgr1term 24061* | The degree of a monomial. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ 𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧↑𝑁))) ⇒ ⊢ ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝑁 ∈ ℕ0) → (deg‘𝐹) = 𝑁) | ||
Theorem | plycn 24062 | A polynomial is a continuous function. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
Theorem | dgr0 24063 | The degree of the zero polynomial is zero. Note: this differs from some other definitions of the degree of the zero polynomial, such as -1, -∞ or undefined. But it is convenient for us to define it this way, so that we have dgrcl 24034, dgreq0 24066 and coeid 24039 without having to special-case zero, although plydivalg 24099 is a little more complicated as a result. (Contributed by Mario Carneiro, 22-Jul-2014.) |
⊢ (deg‘0𝑝) = 0 | ||
Theorem | coeidp 24064 | The coefficients of the identity function. (Contributed by Mario Carneiro, 28-Jul-2014.) |
⊢ (𝐴 ∈ ℕ0 → ((coeff‘Xp)‘𝐴) = if(𝐴 = 1, 1, 0)) | ||
Theorem | dgrid 24065 | The degree of the identity function. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ (deg‘Xp) = 1 | ||
Theorem | dgreq0 24066 | The leading coefficient of a polynomial is nonzero, unless the entire polynomial is zero. (Contributed by Mario Carneiro, 22-Jul-2014.) (Proof shortened by Fan Zheng, 21-Jun-2016.) |
⊢ 𝑁 = (deg‘𝐹) & ⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (𝐹 = 0𝑝 ↔ (𝐴‘𝑁) = 0)) | ||
Theorem | dgrlt 24067 | Two ways to say that the degree of 𝐹 is strictly less than 𝑁. (Contributed by Mario Carneiro, 25-Jul-2014.) |
⊢ 𝑁 = (deg‘𝐹) & ⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑀 ∈ ℕ0) → ((𝐹 = 0𝑝 ∨ 𝑁 < 𝑀) ↔ (𝑁 ≤ 𝑀 ∧ (𝐴‘𝑀) = 0))) | ||
Theorem | dgradd 24068 | The degree of a sum of polynomials is at most the maximum of the degrees. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (deg‘(𝐹 ∘𝑓 + 𝐺)) ≤ if(𝑀 ≤ 𝑁, 𝑁, 𝑀)) | ||
Theorem | dgradd2 24069 | The degree of a sum of polynomials of unequal degrees is the degree of the larger polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆) ∧ 𝑀 < 𝑁) → (deg‘(𝐹 ∘𝑓 + 𝐺)) = 𝑁) | ||
Theorem | dgrmul2 24070 | The degree of a product of polynomials is at most the sum of degrees. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (deg‘(𝐹 ∘𝑓 · 𝐺)) ≤ (𝑀 + 𝑁)) | ||
Theorem | dgrmul 24071 | The degree of a product of nonzero polynomials is the sum of degrees. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ (((𝐹 ∈ (Poly‘𝑆) ∧ 𝐹 ≠ 0𝑝) ∧ (𝐺 ∈ (Poly‘𝑆) ∧ 𝐺 ≠ 0𝑝)) → (deg‘(𝐹 ∘𝑓 · 𝐺)) = (𝑀 + 𝑁)) | ||
Theorem | dgrmulc 24072 | Scalar multiplication by a nonzero constant does not change the degree of a function. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 𝐹 ∈ (Poly‘𝑆)) → (deg‘((ℂ × {𝐴}) ∘𝑓 · 𝐹)) = (deg‘𝐹)) | ||
Theorem | dgrsub 24073 | The degree of a difference of polynomials is at most the maximum of the degrees. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (deg‘(𝐹 ∘𝑓 − 𝐺)) ≤ if(𝑀 ≤ 𝑁, 𝑁, 𝑀)) | ||
Theorem | dgrcolem1 24074* | The degree of a composition of a monomial with a polynomial. (Contributed by Mario Carneiro, 15-Sep-2014.) |
⊢ 𝑁 = (deg‘𝐺) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) ⇒ ⊢ (𝜑 → (deg‘(𝑥 ∈ ℂ ↦ ((𝐺‘𝑥)↑𝑀))) = (𝑀 · 𝑁)) | ||
Theorem | dgrcolem2 24075* | Lemma for dgrco 24076. (Contributed by Mario Carneiro, 15-Sep-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ 𝐴 = (coeff‘𝐹) & ⊢ (𝜑 → 𝐷 ∈ ℕ0) & ⊢ (𝜑 → 𝑀 = (𝐷 + 1)) & ⊢ (𝜑 → ∀𝑓 ∈ (Poly‘ℂ)((deg‘𝑓) ≤ 𝐷 → (deg‘(𝑓 ∘ 𝐺)) = ((deg‘𝑓) · 𝑁))) ⇒ ⊢ (𝜑 → (deg‘(𝐹 ∘ 𝐺)) = (𝑀 · 𝑁)) | ||
Theorem | dgrco 24076 | The degree of a composition of two polynomials is the product of the degrees. (Contributed by Mario Carneiro, 15-Sep-2014.) |
⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) ⇒ ⊢ (𝜑 → (deg‘(𝐹 ∘ 𝐺)) = (𝑀 · 𝑁)) | ||
Theorem | plycjlem 24077* | Lemma for plycj 24078 and coecj 24079. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑁 = (deg‘𝐹) & ⊢ 𝐺 = ((∗ ∘ 𝐹) ∘ ∗) & ⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(((∗ ∘ 𝐴)‘𝑘) · (𝑧↑𝑘)))) | ||
Theorem | plycj 24078* | The double conjugation of a polynomial is a polynomial. (The single conjugation is not because our definition of polynomial includes only holomorphic functions, i.e. no dependence on (∗‘𝑧) independently of 𝑧.) (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑁 = (deg‘𝐹) & ⊢ 𝐺 = ((∗ ∘ 𝐹) ∘ ∗) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝑆) → (∗‘𝑥) ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) ⇒ ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) | ||
Theorem | coecj 24079 | Double conjugation of a polynomial causes the coefficients to be conjugated. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ 𝑁 = (deg‘𝐹) & ⊢ 𝐺 = ((∗ ∘ 𝐹) ∘ ∗) & ⊢ 𝐴 = (coeff‘𝐹) ⇒ ⊢ (𝐹 ∈ (Poly‘𝑆) → (coeff‘𝐺) = (∗ ∘ 𝐴)) | ||
Theorem | plyrecj 24080 | A polynomial with real coefficients distributes under conjugation. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝐹 ∈ (Poly‘ℝ) ∧ 𝐴 ∈ ℂ) → (∗‘(𝐹‘𝐴)) = (𝐹‘(∗‘𝐴))) | ||
Theorem | plymul0or 24081 | Polynomial multiplication has no zero divisors. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → ((𝐹 ∘𝑓 · 𝐺) = 0𝑝 ↔ (𝐹 = 0𝑝 ∨ 𝐺 = 0𝑝))) | ||
Theorem | ofmulrt 24082 | The set of roots of a product is the union of the roots of the terms. (Contributed by Mario Carneiro, 28-Jul-2014.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐹:𝐴⟶ℂ ∧ 𝐺:𝐴⟶ℂ) → (◡(𝐹 ∘𝑓 · 𝐺) “ {0}) = ((◡𝐹 “ {0}) ∪ (◡𝐺 “ {0}))) | ||
Theorem | plyreres 24083 | Real-coefficient polynomials restrict to real functions. (Contributed by Stefan O'Rear, 16-Nov-2014.) |
⊢ (𝐹 ∈ (Poly‘ℝ) → (𝐹 ↾ ℝ):ℝ⟶ℝ) | ||
Theorem | dvply1 24084* | Derivative of a polynomial, explicit sum version. (Contributed by Stefan O'Rear, 13-Nov-2014.) (Revised by Mario Carneiro, 11-Feb-2015.) |
⊢ (𝜑 → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...(𝑁 − 1))((𝐵‘𝑘) · (𝑧↑𝑘)))) & ⊢ (𝜑 → 𝐴:ℕ0⟶ℂ) & ⊢ 𝐵 = (𝑘 ∈ ℕ0 ↦ ((𝑘 + 1) · (𝐴‘(𝑘 + 1)))) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) ⇒ ⊢ (𝜑 → (ℂ D 𝐹) = 𝐺) | ||
Theorem | dvply2g 24085 | The derivative of a polynomial with coefficients in a subring is a polynomial with coefficients in the same ring. (Contributed by Mario Carneiro, 1-Jan-2017.) |
⊢ ((𝑆 ∈ (SubRing‘ℂfld) ∧ 𝐹 ∈ (Poly‘𝑆)) → (ℂ D 𝐹) ∈ (Poly‘𝑆)) | ||
Theorem | dvply2 24086 | The derivative of a polynomial is a polynomial. (Contributed by Stefan O'Rear, 14-Nov-2014.) (Proof shortened by Mario Carneiro, 1-Jan-2017.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → (ℂ D 𝐹) ∈ (Poly‘ℂ)) | ||
Theorem | dvnply2 24087 | Polynomials have polynomials as derivatives of all orders. (Contributed by Mario Carneiro, 1-Jan-2017.) |
⊢ ((𝑆 ∈ (SubRing‘ℂfld) ∧ 𝐹 ∈ (Poly‘𝑆) ∧ 𝑁 ∈ ℕ0) → ((ℂ D𝑛 𝐹)‘𝑁) ∈ (Poly‘𝑆)) | ||
Theorem | dvnply 24088 | Polynomials have polynomials as derivatives of all orders. (Contributed by Stefan O'Rear, 15-Nov-2014.) (Revised by Mario Carneiro, 1-Jan-2017.) |
⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝑁 ∈ ℕ0) → ((ℂ D𝑛 𝐹)‘𝑁) ∈ (Poly‘ℂ)) | ||
Theorem | plycpn 24089 | Polynomials are smooth. (Contributed by Stefan O'Rear, 16-Nov-2014.) (Revised by Mario Carneiro, 11-Feb-2015.) |
⊢ (𝐹 ∈ (Poly‘𝑆) → 𝐹 ∈ ∩ ran (Cn‘ℂ)) | ||
Syntax | cquot 24090 | Extend class notation to include the quotient of a polynomial division. |
class quot | ||
Definition | df-quot 24091* | Define the quotient function on polynomials. This is the 𝑞 of the expression 𝑓 = 𝑔 · 𝑞 + 𝑟 in the division algorithm. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ quot = (𝑓 ∈ (Poly‘ℂ), 𝑔 ∈ ((Poly‘ℂ) ∖ {0𝑝}) ↦ (℩𝑞 ∈ (Poly‘ℂ)[(𝑓 ∘𝑓 − (𝑔 ∘𝑓 · 𝑞)) / 𝑟](𝑟 = 0𝑝 ∨ (deg‘𝑟) < (deg‘𝑔)))) | ||
Theorem | quotval 24092* | Value of the quotient function. (Contributed by Mario Carneiro, 23-Jul-2014.) |
⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) ⇒ ⊢ ((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆) ∧ 𝐺 ≠ 0𝑝) → (𝐹 quot 𝐺) = (℩𝑞 ∈ (Poly‘ℂ)(𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺)))) | ||
Theorem | plydivlem1 24093* | Lemma for plydivalg 24099. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) ⇒ ⊢ (𝜑 → 0 ∈ 𝑆) | ||
Theorem | plydivlem2 24094* | Lemma for plydivalg 24099. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) ⇒ ⊢ ((𝜑 ∧ 𝑞 ∈ (Poly‘𝑆)) → 𝑅 ∈ (Poly‘𝑆)) | ||
Theorem | plydivlem3 24095* | Lemma for plydivex 24097. Base case of induction. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) & ⊢ (𝜑 → (𝐹 = 0𝑝 ∨ ((deg‘𝐹) − (deg‘𝐺)) < 0)) ⇒ ⊢ (𝜑 → ∃𝑞 ∈ (Poly‘𝑆)(𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺))) | ||
Theorem | plydivlem4 24096* | Lemma for plydivex 24097. Induction step. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) & ⊢ (𝜑 → 𝐷 ∈ ℕ0) & ⊢ (𝜑 → (𝑀 − 𝑁) = 𝐷) & ⊢ (𝜑 → 𝐹 ≠ 0𝑝) & ⊢ 𝑈 = (𝑓 ∘𝑓 − (𝐺 ∘𝑓 · 𝑝)) & ⊢ 𝐻 = (𝑧 ∈ ℂ ↦ (((𝐴‘𝑀) / (𝐵‘𝑁)) · (𝑧↑𝐷))) & ⊢ (𝜑 → ∀𝑓 ∈ (Poly‘𝑆)((𝑓 = 0𝑝 ∨ ((deg‘𝑓) − 𝑁) < 𝐷) → ∃𝑝 ∈ (Poly‘𝑆)(𝑈 = 0𝑝 ∨ (deg‘𝑈) < 𝑁))) & ⊢ 𝐴 = (coeff‘𝐹) & ⊢ 𝐵 = (coeff‘𝐺) & ⊢ 𝑀 = (deg‘𝐹) & ⊢ 𝑁 = (deg‘𝐺) ⇒ ⊢ (𝜑 → ∃𝑞 ∈ (Poly‘𝑆)(𝑅 = 0𝑝 ∨ (deg‘𝑅) < 𝑁)) | ||
Theorem | plydivex 24097* | Lemma for plydivalg 24099. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) ⇒ ⊢ (𝜑 → ∃𝑞 ∈ (Poly‘𝑆)(𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺))) | ||
Theorem | plydiveu 24098* | Lemma for plydivalg 24099. (Contributed by Mario Carneiro, 24-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) & ⊢ (𝜑 → 𝑞 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → (𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺))) & ⊢ 𝑇 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑝)) & ⊢ (𝜑 → 𝑝 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → (𝑇 = 0𝑝 ∨ (deg‘𝑇) < (deg‘𝐺))) ⇒ ⊢ (𝜑 → 𝑝 = 𝑞) | ||
Theorem | plydivalg 24099* | The division algorithm on polynomials over a subfield 𝑆 of the complex numbers. If 𝐹 and 𝐺 ≠ 0 are polynomials over 𝑆, then there is a unique quotient polynomial 𝑞 such that the remainder 𝐹 − 𝐺 · 𝑞 is either zero or has degree less than 𝐺. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · 𝑞)) ⇒ ⊢ (𝜑 → ∃!𝑞 ∈ (Poly‘𝑆)(𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺))) | ||
Theorem | quotlem 24100* | Lemma for properties of the polynomial quotient function. (Contributed by Mario Carneiro, 26-Jul-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑥 ≠ 0)) → (1 / 𝑥) ∈ 𝑆) & ⊢ (𝜑 → -1 ∈ 𝑆) & ⊢ (𝜑 → 𝐹 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ∈ (Poly‘𝑆)) & ⊢ (𝜑 → 𝐺 ≠ 0𝑝) & ⊢ 𝑅 = (𝐹 ∘𝑓 − (𝐺 ∘𝑓 · (𝐹 quot 𝐺))) ⇒ ⊢ (𝜑 → ((𝐹 quot 𝐺) ∈ (Poly‘𝑆) ∧ (𝑅 = 0𝑝 ∨ (deg‘𝑅) < (deg‘𝐺)))) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |