HomeHome Intuitionistic Logic Explorer
Theorem List (p. 149 of 156)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 14801-14900   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremhovera 14801* A point at which the hover function is less than a given value. (Contributed by Jim Kingdon, 21-Jul-2025.)
𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < ))       (𝑍 ∈ ℝ → (𝐹‘(𝑍 − 1)) < 𝑍)
 
Theoremhoverb 14802* A point at which the hover function is greater than a given value. (Contributed by Jim Kingdon, 21-Jul-2025.)
𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < ))       (𝑍 ∈ ℝ → 𝑍 < (𝐹‘(𝑍 + 2)))
 
Theoremhoverlt1 14803* The hover function evaluated at a point less than one. (Contributed by Jim Kingdon, 22-Jul-2025.)
𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < ))       ((𝐶 ∈ ℝ ∧ 𝐶 < 1) → (𝐹𝐶) ≤ 0)
 
Theoremhovergt0 14804* The hover function evaluated at a point greater than zero. (Contributed by Jim Kingdon, 22-Jul-2025.)
𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < ))       ((𝐶 ∈ ℝ ∧ 0 < 𝐶) → 0 ≤ (𝐹𝐶))
 
Theoremivthdichlem 14805* Lemma for ivthdich 14807. The result, with a few notational conveniences. (Contributed by Jim Kingdon, 22-Jul-2025.)
𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < ))    &   (𝜑𝑍 ∈ ℝ)    &   (𝜑 → ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓𝑎) < 0 ∧ 0 < (𝑓𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥𝑥 < 𝑏 ∧ (𝑓𝑥) = 0))))       (𝜑 → (𝑍 ≤ 0 ∨ 0 ≤ 𝑍))
 
Theoremdich0 14806* Real number dichotomy stated in terms of two real numbers or a real number and zero. (Contributed by Jim Kingdon, 22-Jul-2025.)
(∀𝑧 ∈ ℝ (𝑧 ≤ 0 ∨ 0 ≤ 𝑧) ↔ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥𝑦𝑦𝑥))
 
Theoremivthdich 14807* The intermediate value theorem implies real number dichotomy. Because real number dichotomy (also known as analytic LLPO) is a constructive taboo, this means we will be unable to prove the intermediate value theorem as stated here (although versions with additional conditions, such as ivthinc 14797 for strictly monotonic functions, can be proved).

The proof is via a function which we call the hover function and which is also described in Section 5.1 of [Bauer], p. 493. Consider any real number 𝑧. We want to show that 𝑧 ≤ 0 ∨ 0 ≤ 𝑧. Because of hovercncf 14800, hovera 14801, and hoverb 14802, we are able to apply the intermediate value theorem to get a value 𝑐 such that the hover function at 𝑐 equals 𝑧. By axltwlin 8087, 𝑐 < 1 or 0 < 𝑐, and that leads to 𝑧 ≤ 0 by hoverlt1 14803 or 0 ≤ 𝑧 by hovergt0 14804. (Contributed by Jim Kingdon and Mario Carneiro, 22-Jul-2025.)

(∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓𝑎) < 0 ∧ 0 < (𝑓𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥𝑥 < 𝑏 ∧ (𝑓𝑥) = 0))) → ∀𝑟 ∈ ℝ ∀𝑠 ∈ ℝ (𝑟𝑠𝑠𝑟))
 
10.2  Derivatives
 
10.2.1  Real and complex differentiation
 
10.2.1.1  Derivatives of functions of one complex or real variable
 
Syntaxclimc 14808 The limit operator.
class lim
 
Syntaxcdv 14809 The derivative operator.
class D
 
Definitiondf-limced 14810* Define the set of limits of a complex function at a point. Under normal circumstances, this will be a singleton or empty, depending on whether the limit exists. (Contributed by Mario Carneiro, 24-Dec-2016.) (Revised by Jim Kingdon, 3-Jun-2023.)
lim = (𝑓 ∈ (ℂ ↑pm ℂ), 𝑥 ∈ ℂ ↦ {𝑦 ∈ ℂ ∣ ((𝑓:dom 𝑓⟶ℂ ∧ dom 𝑓 ⊆ ℂ) ∧ (𝑥 ∈ ℂ ∧ ∀𝑒 ∈ ℝ+𝑑 ∈ ℝ+𝑧 ∈ dom 𝑓((𝑧 # 𝑥 ∧ (abs‘(𝑧𝑥)) < 𝑑) → (abs‘((𝑓𝑧) − 𝑦)) < 𝑒)))})
 
Definitiondf-dvap 14811* Define the derivative operator. This acts on functions to produce a function that is defined where the original function is differentiable, with value the derivative of the function at these points. The set 𝑠 here is the ambient topological space under which we are evaluating the continuity of the difference quotient. Although the definition is valid for any subset of and is well-behaved when 𝑠 contains no isolated points, we will restrict our attention to the cases 𝑠 = ℝ or 𝑠 = ℂ for the majority of the development, these corresponding respectively to real and complex differentiation. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 25-Jun-2023.)
D = (𝑠 ∈ 𝒫 ℂ, 𝑓 ∈ (ℂ ↑pm 𝑠) ↦ 𝑥 ∈ ((int‘((MetOpen‘(abs ∘ − )) ↾t 𝑠))‘dom 𝑓)({𝑥} × ((𝑧 ∈ {𝑤 ∈ dom 𝑓𝑤 # 𝑥} ↦ (((𝑓𝑧) − (𝑓𝑥)) / (𝑧𝑥))) lim 𝑥)))
 
Theoremlimcrcl 14812 Reverse closure for the limit operator. (Contributed by Mario Carneiro, 28-Dec-2016.)
(𝐶 ∈ (𝐹 lim 𝐵) → (𝐹:dom 𝐹⟶ℂ ∧ dom 𝐹 ⊆ ℂ ∧ 𝐵 ∈ ℂ))
 
Theoremlimccl 14813 Closure of the limit operator. (Contributed by Mario Carneiro, 25-Dec-2016.)
(𝐹 lim 𝐵) ⊆ ℂ
 
Theoremellimc3apf 14814* Write the epsilon-delta definition of a limit. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 4-Nov-2023.)
(𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   𝑧𝐹       (𝜑 → (𝐶 ∈ (𝐹 lim 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+𝑦 ∈ ℝ+𝑧𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧𝐵)) < 𝑦) → (abs‘((𝐹𝑧) − 𝐶)) < 𝑥))))
 
Theoremellimc3ap 14815* Write the epsilon-delta definition of a limit. (Contributed by Mario Carneiro, 28-Dec-2016.) Use apartness. (Revised by Jim Kingdon, 3-Jun-2023.)
(𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐵 ∈ ℂ)       (𝜑 → (𝐶 ∈ (𝐹 lim 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+𝑦 ∈ ℝ+𝑧𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧𝐵)) < 𝑦) → (abs‘((𝐹𝑧) − 𝐶)) < 𝑥))))
 
Theoremlimcdifap 14816* It suffices to consider functions which are not defined at 𝐵 to define the limit of a function. In particular, the value of the original function 𝐹 at 𝐵 does not affect the limit of 𝐹. (Contributed by Mario Carneiro, 25-Dec-2016.) (Revised by Jim Kingdon, 3-Jun-2023.)
(𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴 ⊆ ℂ)       (𝜑 → (𝐹 lim 𝐵) = ((𝐹 ↾ {𝑥𝐴𝑥 # 𝐵}) lim 𝐵))
 
Theoremlimcmpted 14817* Express the limit operator for a function defined by a mapping, via epsilon-delta. (Contributed by Jim Kingdon, 3-Nov-2023.)
(𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   ((𝜑𝑧𝐴) → 𝐷 ∈ ℂ)       (𝜑 → (𝐶 ∈ ((𝑧𝐴𝐷) lim 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+𝑦 ∈ ℝ+𝑧𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧𝐵)) < 𝑦) → (abs‘(𝐷𝐶)) < 𝑥))))
 
Theoremlimcimolemlt 14818* Lemma for limcimo 14819. (Contributed by Jim Kingdon, 3-Jul-2023.)
(𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐵𝐶)    &   (𝜑𝐵𝑆)    &   (𝜑𝐶 ∈ (𝐾t 𝑆))    &   (𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑 → {𝑞𝐶𝑞 # 𝐵} ⊆ 𝐴)    &   𝐾 = (MetOpen‘(abs ∘ − ))    &   (𝜑𝐷 ∈ ℝ+)    &   (𝜑𝑋 ∈ (𝐹 lim 𝐵))    &   (𝜑𝑌 ∈ (𝐹 lim 𝐵))    &   (𝜑 → ∀𝑧𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧𝐵)) < 𝐷) → (abs‘((𝐹𝑧) − 𝑋)) < ((abs‘(𝑋𝑌)) / 2)))    &   (𝜑𝐺 ∈ ℝ+)    &   (𝜑 → ∀𝑤𝐴 ((𝑤 # 𝐵 ∧ (abs‘(𝑤𝐵)) < 𝐺) → (abs‘((𝐹𝑤) − 𝑌)) < ((abs‘(𝑋𝑌)) / 2)))       (𝜑 → (abs‘(𝑋𝑌)) < (abs‘(𝑋𝑌)))
 
Theoremlimcimo 14819* Conditions which ensure there is at most one limit value of 𝐹 at 𝐵. (Contributed by Mario Carneiro, 25-Dec-2016.) (Revised by Jim Kingdon, 8-Jul-2023.)
(𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐵 ∈ ℂ)    &   (𝜑𝐵𝐶)    &   (𝜑𝐵𝑆)    &   (𝜑𝐶 ∈ (𝐾t 𝑆))    &   (𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑 → {𝑞𝐶𝑞 # 𝐵} ⊆ 𝐴)    &   𝐾 = (MetOpen‘(abs ∘ − ))       (𝜑 → ∃*𝑥 𝑥 ∈ (𝐹 lim 𝐵))
 
Theoremlimcresi 14820 Any limit of 𝐹 is also a limit of the restriction of 𝐹. (Contributed by Mario Carneiro, 28-Dec-2016.)
(𝐹 lim 𝐵) ⊆ ((𝐹𝐶) lim 𝐵)
 
Theoremcnplimcim 14821 If a function is continuous at 𝐵, its limit at 𝐵 equals the value of the function there. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 14-Jun-2023.)
𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = (𝐾t 𝐴)       ((𝐴 ⊆ ℂ ∧ 𝐵𝐴) → (𝐹 ∈ ((𝐽 CnP 𝐾)‘𝐵) → (𝐹:𝐴⟶ℂ ∧ (𝐹𝐵) ∈ (𝐹 lim 𝐵))))
 
Theoremcnplimclemle 14822 Lemma for cnplimccntop 14824. Satisfying the epsilon condition for continuity. (Contributed by Mario Carneiro and Jim Kingdon, 17-Nov-2023.)
𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = (𝐾t 𝐴)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐵𝐴)    &   (𝜑 → (𝐹𝐵) ∈ (𝐹 lim 𝐵))    &   (𝜑𝐸 ∈ ℝ+)    &   (𝜑𝐷 ∈ ℝ+)    &   (𝜑𝑍𝐴)    &   ((𝜑𝑍 # 𝐵 ∧ (abs‘(𝑍𝐵)) < 𝐷) → (abs‘((𝐹𝑍) − (𝐹𝐵))) < (𝐸 / 2))    &   (𝜑 → (abs‘(𝑍𝐵)) < 𝐷)       (𝜑 → (abs‘((𝐹𝑍) − (𝐹𝐵))) < 𝐸)
 
Theoremcnplimclemr 14823 Lemma for cnplimccntop 14824. The reverse direction. (Contributed by Mario Carneiro and Jim Kingdon, 17-Nov-2023.)
𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = (𝐾t 𝐴)    &   (𝜑𝐴 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐵𝐴)    &   (𝜑 → (𝐹𝐵) ∈ (𝐹 lim 𝐵))       (𝜑𝐹 ∈ ((𝐽 CnP 𝐾)‘𝐵))
 
Theoremcnplimccntop 14824 A function is continuous at 𝐵 iff its limit at 𝐵 equals the value of the function there. (Contributed by Mario Carneiro, 28-Dec-2016.)
𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = (𝐾t 𝐴)       ((𝐴 ⊆ ℂ ∧ 𝐵𝐴) → (𝐹 ∈ ((𝐽 CnP 𝐾)‘𝐵) ↔ (𝐹:𝐴⟶ℂ ∧ (𝐹𝐵) ∈ (𝐹 lim 𝐵))))
 
Theoremcnlimcim 14825* If 𝐹 is a continuous function, the limit of the function at each point equals the value of the function. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 16-Jun-2023.)
(𝐴 ⊆ ℂ → (𝐹 ∈ (𝐴cn→ℂ) → (𝐹:𝐴⟶ℂ ∧ ∀𝑥𝐴 (𝐹𝑥) ∈ (𝐹 lim 𝑥))))
 
Theoremcnlimc 14826* 𝐹 is a continuous function iff the limit of the function at each point equals the value of the function. (Contributed by Mario Carneiro, 28-Dec-2016.)
(𝐴 ⊆ ℂ → (𝐹 ∈ (𝐴cn→ℂ) ↔ (𝐹:𝐴⟶ℂ ∧ ∀𝑥𝐴 (𝐹𝑥) ∈ (𝐹 lim 𝑥))))
 
Theoremcnlimci 14827 If 𝐹 is a continuous function, then the limit of the function at any point equals its value. (Contributed by Mario Carneiro, 28-Dec-2016.)
(𝜑𝐹 ∈ (𝐴cn𝐷))    &   (𝜑𝐵𝐴)       (𝜑 → (𝐹𝐵) ∈ (𝐹 lim 𝐵))
 
Theoremcnmptlimc 14828* If 𝐹 is a continuous function, then the limit of the function at any point equals its value. (Contributed by Mario Carneiro, 28-Dec-2016.)
(𝜑 → (𝑥𝐴𝑋) ∈ (𝐴cn𝐷))    &   (𝜑𝐵𝐴)    &   (𝑥 = 𝐵𝑋 = 𝑌)       (𝜑𝑌 ∈ ((𝑥𝐴𝑋) lim 𝐵))
 
Theoremlimccnpcntop 14829 If the limit of 𝐹 at 𝐵 is 𝐶 and 𝐺 is continuous at 𝐶, then the limit of 𝐺𝐹 at 𝐵 is 𝐺(𝐶). (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 18-Jun-2023.)
(𝜑𝐹:𝐴𝐷)    &   (𝜑𝐷 ⊆ ℂ)    &   𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = (𝐾t 𝐷)    &   (𝜑𝐶 ∈ (𝐹 lim 𝐵))    &   (𝜑𝐺 ∈ ((𝐽 CnP 𝐾)‘𝐶))       (𝜑 → (𝐺𝐶) ∈ ((𝐺𝐹) lim 𝐵))
 
Theoremlimccnp2lem 14830* Lemma for limccnp2cntop 14831. This is most of the result, expressed in epsilon-delta form, with a large number of hypotheses so that lengthy expressions do not need to be repeated. (Contributed by Jim Kingdon, 9-Nov-2023.)
((𝜑𝑥𝐴) → 𝑅𝑋)    &   ((𝜑𝑥𝐴) → 𝑆𝑌)    &   (𝜑𝑋 ⊆ ℂ)    &   (𝜑𝑌 ⊆ ℂ)    &   𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = ((𝐾 ×t 𝐾) ↾t (𝑋 × 𝑌))    &   (𝜑𝐶 ∈ ((𝑥𝐴𝑅) lim 𝐵))    &   (𝜑𝐷 ∈ ((𝑥𝐴𝑆) lim 𝐵))    &   (𝜑𝐻 ∈ ((𝐽 CnP 𝐾)‘⟨𝐶, 𝐷⟩))    &   𝑥𝜑    &   (𝜑𝐸 ∈ ℝ+)    &   (𝜑𝐿 ∈ ℝ+)    &   (𝜑 → ∀𝑟𝑋𝑠𝑌 (((𝐶((abs ∘ − ) ↾ (𝑋 × 𝑋))𝑟) < 𝐿 ∧ (𝐷((abs ∘ − ) ↾ (𝑌 × 𝑌))𝑠) < 𝐿) → ((𝐶𝐻𝐷)(abs ∘ − )(𝑟𝐻𝑠)) < 𝐸))    &   (𝜑𝐹 ∈ ℝ+)    &   (𝜑 → ∀𝑥𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥𝐵)) < 𝐹) → (abs‘(𝑅𝐶)) < 𝐿))    &   (𝜑𝐺 ∈ ℝ+)    &   (𝜑 → ∀𝑥𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥𝐵)) < 𝐺) → (abs‘(𝑆𝐷)) < 𝐿))       (𝜑 → ∃𝑑 ∈ ℝ+𝑥𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥𝐵)) < 𝑑) → (abs‘((𝑅𝐻𝑆) − (𝐶𝐻𝐷))) < 𝐸))
 
Theoremlimccnp2cntop 14831* The image of a convergent sequence under a continuous map is convergent to the image of the original point. Binary operation version. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 14-Nov-2023.)
((𝜑𝑥𝐴) → 𝑅𝑋)    &   ((𝜑𝑥𝐴) → 𝑆𝑌)    &   (𝜑𝑋 ⊆ ℂ)    &   (𝜑𝑌 ⊆ ℂ)    &   𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐽 = ((𝐾 ×t 𝐾) ↾t (𝑋 × 𝑌))    &   (𝜑𝐶 ∈ ((𝑥𝐴𝑅) lim 𝐵))    &   (𝜑𝐷 ∈ ((𝑥𝐴𝑆) lim 𝐵))    &   (𝜑𝐻 ∈ ((𝐽 CnP 𝐾)‘⟨𝐶, 𝐷⟩))       (𝜑 → (𝐶𝐻𝐷) ∈ ((𝑥𝐴 ↦ (𝑅𝐻𝑆)) lim 𝐵))
 
Theoremlimccoap 14832* Composition of two limits. This theorem is only usable in the case where 𝑥 # 𝑋 implies R(x) # 𝐶 so it is less general than might appear at first. (Contributed by Mario Carneiro, 29-Dec-2016.) (Revised by Jim Kingdon, 18-Dec-2023.)
((𝜑𝑥 ∈ {𝑤𝐴𝑤 # 𝑋}) → 𝑅 ∈ {𝑤𝐵𝑤 # 𝐶})    &   ((𝜑𝑦 ∈ {𝑤𝐵𝑤 # 𝐶}) → 𝑆 ∈ ℂ)    &   (𝜑𝐶 ∈ ((𝑥 ∈ {𝑤𝐴𝑤 # 𝑋} ↦ 𝑅) lim 𝑋))    &   (𝜑𝐷 ∈ ((𝑦 ∈ {𝑤𝐵𝑤 # 𝐶} ↦ 𝑆) lim 𝐶))    &   (𝑦 = 𝑅𝑆 = 𝑇)       (𝜑𝐷 ∈ ((𝑥 ∈ {𝑤𝐴𝑤 # 𝑋} ↦ 𝑇) lim 𝑋))
 
Theoremreldvg 14833 The derivative function is a relation. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 25-Jun-2023.)
((𝑆 ⊆ ℂ ∧ 𝐹 ∈ (ℂ ↑pm 𝑆)) → Rel (𝑆 D 𝐹))
 
Theoremdvlemap 14834* Closure for a difference quotient. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 27-Jun-2023.)
(𝜑𝐹:𝐷⟶ℂ)    &   (𝜑𝐷 ⊆ ℂ)    &   (𝜑𝐵𝐷)       ((𝜑𝐴 ∈ {𝑤𝐷𝑤 # 𝐵}) → (((𝐹𝐴) − (𝐹𝐵)) / (𝐴𝐵)) ∈ ℂ)
 
Theoremdvfvalap 14835* Value and set bounds on the derivative operator. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 27-Jun-2023.)
𝑇 = (𝐾t 𝑆)    &   𝐾 = (MetOpen‘(abs ∘ − ))       ((𝑆 ⊆ ℂ ∧ 𝐹:𝐴⟶ℂ ∧ 𝐴𝑆) → ((𝑆 D 𝐹) = 𝑥 ∈ ((int‘𝑇)‘𝐴)({𝑥} × ((𝑧 ∈ {𝑤𝐴𝑤 # 𝑥} ↦ (((𝐹𝑧) − (𝐹𝑥)) / (𝑧𝑥))) lim 𝑥)) ∧ (𝑆 D 𝐹) ⊆ (((int‘𝑇)‘𝐴) × ℂ)))
 
Theoremeldvap 14836* The differentiable predicate. A function 𝐹 is differentiable at 𝐵 with derivative 𝐶 iff 𝐹 is defined in a neighborhood of 𝐵 and the difference quotient has limit 𝐶 at 𝐵. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 27-Jun-2023.)
𝑇 = (𝐾t 𝑆)    &   𝐾 = (MetOpen‘(abs ∘ − ))    &   𝐺 = (𝑧 ∈ {𝑤𝐴𝑤 # 𝐵} ↦ (((𝐹𝑧) − (𝐹𝐵)) / (𝑧𝐵)))    &   (𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴𝑆)       (𝜑 → (𝐵(𝑆 D 𝐹)𝐶 ↔ (𝐵 ∈ ((int‘𝑇)‘𝐴) ∧ 𝐶 ∈ (𝐺 lim 𝐵))))
 
Theoremdvcl 14837 The derivative function takes values in the complex numbers. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Mario Carneiro, 9-Feb-2015.)
(𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴𝑆)       ((𝜑𝐵(𝑆 D 𝐹)𝐶) → 𝐶 ∈ ℂ)
 
Theoremdvbssntrcntop 14838 The set of differentiable points is a subset of the interior of the domain of the function. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 27-Jun-2023.)
(𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴𝑆)    &   𝐽 = (𝐾t 𝑆)    &   𝐾 = (MetOpen‘(abs ∘ − ))       (𝜑 → dom (𝑆 D 𝐹) ⊆ ((int‘𝐽)‘𝐴))
 
Theoremdvbss 14839 The set of differentiable points is a subset of the domain of the function. (Contributed by Mario Carneiro, 6-Aug-2014.) (Revised by Mario Carneiro, 9-Feb-2015.)
(𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐹:𝐴⟶ℂ)    &   (𝜑𝐴𝑆)       (𝜑 → dom (𝑆 D 𝐹) ⊆ 𝐴)
 
Theoremdvbsssg 14840 The set of differentiable points is a subset of the ambient topology. (Contributed by Mario Carneiro, 18-Mar-2015.) (Revised by Jim Kingdon, 28-Jun-2023.)
((𝑆 ⊆ ℂ ∧ 𝐹 ∈ (ℂ ↑pm 𝑆)) → dom (𝑆 D 𝐹) ⊆ 𝑆)
 
Theoremrecnprss 14841 Both and are subsets of . (Contributed by Mario Carneiro, 10-Feb-2015.)
(𝑆 ∈ {ℝ, ℂ} → 𝑆 ⊆ ℂ)
 
Theoremdvfgg 14842 Explicitly write out the functionality condition on derivative for 𝑆 = ℝ and . (Contributed by Mario Carneiro, 9-Feb-2015.) (Revised by Jim Kingdon, 28-Jun-2023.)
((𝑆 ∈ {ℝ, ℂ} ∧ 𝐹 ∈ (ℂ ↑pm 𝑆)) → (𝑆 D 𝐹):dom (𝑆 D 𝐹)⟶ℂ)
 
Theoremdvfpm 14843 The derivative is a function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 28-Jul-2023.)
(𝐹 ∈ (ℂ ↑pm ℝ) → (ℝ D 𝐹):dom (ℝ D 𝐹)⟶ℂ)
 
Theoremdvfcnpm 14844 The derivative is a function. (Contributed by Mario Carneiro, 9-Feb-2015.) (Revised by Jim Kingdon, 28-Jul-2023.)
(𝐹 ∈ (ℂ ↑pm ℂ) → (ℂ D 𝐹):dom (ℂ D 𝐹)⟶ℂ)
 
Theoremdvidlemap 14845* Lemma for dvid 14847 and dvconst 14846. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.)
(𝜑𝐹:ℂ⟶ℂ)    &   ((𝜑 ∧ (𝑥 ∈ ℂ ∧ 𝑧 ∈ ℂ ∧ 𝑧 # 𝑥)) → (((𝐹𝑧) − (𝐹𝑥)) / (𝑧𝑥)) = 𝐵)    &   𝐵 ∈ ℂ       (𝜑 → (ℂ D 𝐹) = (ℂ × {𝐵}))
 
Theoremdvconst 14846 Derivative of a constant function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.)
(𝐴 ∈ ℂ → (ℂ D (ℂ × {𝐴})) = (ℂ × {0}))
 
Theoremdvid 14847 Derivative of the identity function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.)
(ℂ D ( I ↾ ℂ)) = (ℂ × {1})
 
Theoremdvcnp2cntop 14848 A function is continuous at each point for which it is differentiable. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Mario Carneiro, 28-Dec-2016.)
𝐽 = (𝐾t 𝐴)    &   𝐾 = (MetOpen‘(abs ∘ − ))       (((𝑆 ⊆ ℂ ∧ 𝐹:𝐴⟶ℂ ∧ 𝐴𝑆) ∧ 𝐵 ∈ dom (𝑆 D 𝐹)) → 𝐹 ∈ ((𝐽 CnP 𝐾)‘𝐵))
 
Theoremdvcn 14849 A differentiable function is continuous. (Contributed by Mario Carneiro, 7-Sep-2014.) (Revised by Mario Carneiro, 7-Sep-2015.)
(((𝑆 ⊆ ℂ ∧ 𝐹:𝐴⟶ℂ ∧ 𝐴𝑆) ∧ dom (𝑆 D 𝐹) = 𝐴) → 𝐹 ∈ (𝐴cn→ℂ))
 
Theoremdvaddxxbr 14850 The sum rule for derivatives at a point. That is, if the derivative of 𝐹 at 𝐶 is 𝐾 and the derivative of 𝐺 at 𝐶 is 𝐿, then the derivative of the pointwise sum of those two functions at 𝐶 is 𝐾 + 𝐿. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 25-Nov-2023.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋𝑆)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐶(𝑆 D 𝐹)𝐾)    &   (𝜑𝐶(𝑆 D 𝐺)𝐿)    &   𝐽 = (MetOpen‘(abs ∘ − ))       (𝜑𝐶(𝑆 D (𝐹𝑓 + 𝐺))(𝐾 + 𝐿))
 
Theoremdvmulxxbr 14851 The product rule for derivatives at a point. For the (simpler but more limited) function version, see dvmulxx 14853. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 1-Dec-2023.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋𝑆)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑𝑆 ⊆ ℂ)    &   (𝜑𝐶(𝑆 D 𝐹)𝐾)    &   (𝜑𝐶(𝑆 D 𝐺)𝐿)    &   𝐽 = (MetOpen‘(abs ∘ − ))       (𝜑𝐶(𝑆 D (𝐹𝑓 · 𝐺))((𝐾 · (𝐺𝐶)) + (𝐿 · (𝐹𝐶))))
 
Theoremdvaddxx 14852 The sum rule for derivatives at a point. For the (more general) relation version, see dvaddxxbr 14850. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 25-Nov-2023.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋𝑆)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑𝐶 ∈ dom (𝑆 D 𝐹))    &   (𝜑𝐶 ∈ dom (𝑆 D 𝐺))       (𝜑 → ((𝑆 D (𝐹𝑓 + 𝐺))‘𝐶) = (((𝑆 D 𝐹)‘𝐶) + ((𝑆 D 𝐺)‘𝐶)))
 
Theoremdvmulxx 14853 The product rule for derivatives at a point. For the (more general) relation version, see dvmulxxbr 14851. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 2-Dec-2023.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋𝑆)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑𝐶 ∈ dom (𝑆 D 𝐹))    &   (𝜑𝐶 ∈ dom (𝑆 D 𝐺))       (𝜑 → ((𝑆 D (𝐹𝑓 · 𝐺))‘𝐶) = ((((𝑆 D 𝐹)‘𝐶) · (𝐺𝐶)) + (((𝑆 D 𝐺)‘𝐶) · (𝐹𝐶))))
 
Theoremdviaddf 14854 The sum rule for everywhere-differentiable functions. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
(𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑𝑋𝑆)    &   (𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑 → dom (𝑆 D 𝐹) = 𝑋)    &   (𝜑 → dom (𝑆 D 𝐺) = 𝑋)       (𝜑 → (𝑆 D (𝐹𝑓 + 𝐺)) = ((𝑆 D 𝐹) ∘𝑓 + (𝑆 D 𝐺)))
 
Theoremdvimulf 14855 The product rule for everywhere-differentiable functions. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
(𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑𝑋𝑆)    &   (𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝐺:𝑋⟶ℂ)    &   (𝜑 → dom (𝑆 D 𝐹) = 𝑋)    &   (𝜑 → dom (𝑆 D 𝐺) = 𝑋)       (𝜑 → (𝑆 D (𝐹𝑓 · 𝐺)) = (((𝑆 D 𝐹) ∘𝑓 · 𝐺) ∘𝑓 + ((𝑆 D 𝐺) ∘𝑓 · 𝐹)))
 
Theoremdvcoapbr 14856* The chain rule for derivatives at a point. The 𝑢 # 𝐶 → (𝐺𝑢) # (𝐺𝐶) hypothesis constrains what functions work for 𝐺. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 21-Dec-2023.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋𝑆)    &   (𝜑𝐺:𝑌𝑋)    &   (𝜑𝑌𝑇)    &   (𝜑 → ∀𝑢𝑌 (𝑢 # 𝐶 → (𝐺𝑢) # (𝐺𝐶)))    &   (𝜑𝑆 ⊆ ℂ)    &   (𝜑𝑇 ⊆ ℂ)    &   (𝜑 → (𝐺𝐶)(𝑆 D 𝐹)𝐾)    &   (𝜑𝐶(𝑇 D 𝐺)𝐿)    &   𝐽 = (MetOpen‘(abs ∘ − ))       (𝜑𝐶(𝑇 D (𝐹𝐺))(𝐾 · 𝐿))
 
Theoremdvcjbr 14857 The derivative of the conjugate of a function. For the (simpler but more limited) function version, see dvcj 14858. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
(𝜑𝐹:𝑋⟶ℂ)    &   (𝜑𝑋 ⊆ ℝ)    &   (𝜑𝐶 ∈ dom (ℝ D 𝐹))       (𝜑𝐶(ℝ D (∗ ∘ 𝐹))(∗‘((ℝ D 𝐹)‘𝐶)))
 
Theoremdvcj 14858 The derivative of the conjugate of a function. For the (more general) relation version, see dvcjbr 14857. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
((𝐹:𝑋⟶ℂ ∧ 𝑋 ⊆ ℝ) → (ℝ D (∗ ∘ 𝐹)) = (∗ ∘ (ℝ D 𝐹)))
 
Theoremdvfre 14859 The derivative of a real function is real. (Contributed by Mario Carneiro, 1-Sep-2014.)
((𝐹:𝐴⟶ℝ ∧ 𝐴 ⊆ ℝ) → (ℝ D 𝐹):dom (ℝ D 𝐹)⟶ℝ)
 
Theoremdvexp 14860* Derivative of a power function. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
(𝑁 ∈ ℕ → (ℂ D (𝑥 ∈ ℂ ↦ (𝑥𝑁))) = (𝑥 ∈ ℂ ↦ (𝑁 · (𝑥↑(𝑁 − 1)))))
 
Theoremdvexp2 14861* Derivative of an exponential, possibly zero power. (Contributed by Stefan O'Rear, 13-Nov-2014.) (Revised by Mario Carneiro, 10-Feb-2015.)
(𝑁 ∈ ℕ0 → (ℂ D (𝑥 ∈ ℂ ↦ (𝑥𝑁))) = (𝑥 ∈ ℂ ↦ if(𝑁 = 0, 0, (𝑁 · (𝑥↑(𝑁 − 1))))))
 
Theoremdvrecap 14862* Derivative of the reciprocal function. (Contributed by Mario Carneiro, 25-Feb-2015.) (Revised by Mario Carneiro, 28-Dec-2016.)
(𝐴 ∈ ℂ → (ℂ D (𝑥 ∈ {𝑤 ∈ ℂ ∣ 𝑤 # 0} ↦ (𝐴 / 𝑥))) = (𝑥 ∈ {𝑤 ∈ ℂ ∣ 𝑤 # 0} ↦ -(𝐴 / (𝑥↑2))))
 
Theoremdvmptidcn 14863 Function-builder for derivative: derivative of the identity. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 30-Dec-2023.)
(ℂ D (𝑥 ∈ ℂ ↦ 𝑥)) = (𝑥 ∈ ℂ ↦ 1)
 
Theoremdvmptccn 14864* Function-builder for derivative: derivative of a constant. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 30-Dec-2023.)
(𝜑𝐴 ∈ ℂ)       (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 0))
 
Theoremdvmptclx 14865* Closure lemma for dvmptmulx 14867 and other related theorems. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Mario Carneiro, 11-Feb-2015.)
(𝜑𝑆 ∈ {ℝ, ℂ})    &   ((𝜑𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐵𝑉)    &   (𝜑 → (𝑆 D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))    &   (𝜑𝑋𝑆)       ((𝜑𝑥𝑋) → 𝐵 ∈ ℂ)
 
Theoremdvmptaddx 14866* Function-builder for derivative, addition rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Mario Carneiro, 11-Feb-2015.)
(𝜑𝑆 ∈ {ℝ, ℂ})    &   ((𝜑𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐵𝑉)    &   (𝜑 → (𝑆 D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))    &   (𝜑𝑋𝑆)    &   ((𝜑𝑥𝑋) → 𝐶 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐷𝑊)    &   (𝜑 → (𝑆 D (𝑥𝑋𝐶)) = (𝑥𝑋𝐷))       (𝜑 → (𝑆 D (𝑥𝑋 ↦ (𝐴 + 𝐶))) = (𝑥𝑋 ↦ (𝐵 + 𝐷)))
 
Theoremdvmptmulx 14867* Function-builder for derivative, product rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Mario Carneiro, 11-Feb-2015.)
(𝜑𝑆 ∈ {ℝ, ℂ})    &   ((𝜑𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐵𝑉)    &   (𝜑 → (𝑆 D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))    &   (𝜑𝑋𝑆)    &   ((𝜑𝑥𝑋) → 𝐶 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐷𝑊)    &   (𝜑 → (𝑆 D (𝑥𝑋𝐶)) = (𝑥𝑋𝐷))       (𝜑 → (𝑆 D (𝑥𝑋 ↦ (𝐴 · 𝐶))) = (𝑥𝑋 ↦ ((𝐵 · 𝐶) + (𝐷 · 𝐴))))
 
Theoremdvmptcmulcn 14868* Function-builder for derivative, product rule for constant multiplier. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.)
((𝜑𝑥 ∈ ℂ) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐵𝑉)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵))    &   (𝜑𝐶 ∈ ℂ)       (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ (𝐶 · 𝐴))) = (𝑥 ∈ ℂ ↦ (𝐶 · 𝐵)))
 
Theoremdvmptnegcn 14869* Function-builder for derivative, product rule for negatives. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.)
((𝜑𝑥 ∈ ℂ) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐵𝑉)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵))       (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ -𝐴)) = (𝑥 ∈ ℂ ↦ -𝐵))
 
Theoremdvmptsubcn 14870* Function-builder for derivative, subtraction rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.)
((𝜑𝑥 ∈ ℂ) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐵𝑉)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵))    &   ((𝜑𝑥 ∈ ℂ) → 𝐶 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐷𝑊)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐶)) = (𝑥 ∈ ℂ ↦ 𝐷))       (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ (𝐴𝐶))) = (𝑥 ∈ ℂ ↦ (𝐵𝐷)))
 
Theoremdvmptcjx 14871* Function-builder for derivative, conjugate rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 24-May-2024.)
((𝜑𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐵𝑉)    &   (𝜑 → (ℝ D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))    &   (𝜑𝑋 ⊆ ℝ)       (𝜑 → (ℝ D (𝑥𝑋 ↦ (∗‘𝐴))) = (𝑥𝑋 ↦ (∗‘𝐵)))
 
Theoremdveflem 14872 Derivative of the exponential function at 0. The key step in the proof is eftlub 11833, to show that abs(exp(𝑥) − 1 − 𝑥) ≤ abs(𝑥)↑2 · (3 / 4). (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Mario Carneiro, 28-Dec-2016.)
0(ℂ D exp)1
 
Theoremdvef 14873 Derivative of the exponential function. (Contributed by Mario Carneiro, 9-Aug-2014.) (Proof shortened by Mario Carneiro, 10-Feb-2015.)
(ℂ D exp) = exp
 
PART 11  BASIC REAL AND COMPLEX FUNCTIONS
 
11.1  Polynomials
 
11.1.1  Elementary properties of complex polynomials
 
Syntaxcply 14874 Extend class notation to include the set of complex polynomials.
class Poly
 
Syntaxcidp 14875 Extend class notation to include the identity polynomial.
class Xp
 
Definitiondf-ply 14876* Define the set of polynomials on the complex numbers with coefficients in the given subset. (Contributed by Mario Carneiro, 17-Jul-2014.)
Poly = (𝑥 ∈ 𝒫 ℂ ↦ {𝑓 ∣ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑥 ∪ {0}) ↑𝑚0)𝑓 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘)))})
 
Definitiondf-idp 14877 Define the identity polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
Xp = ( I ↾ ℂ)
 
Theoremplyval 14878* Value of the polynomial set function. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝑆 ⊆ ℂ → (Poly‘𝑆) = {𝑓 ∣ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑆 ∪ {0}) ↑𝑚0)𝑓 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘)))})
 
Theoremplybss 14879 Reverse closure of the parameter 𝑆 of the polynomial set function. (Contributed by Mario Carneiro, 22-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) → 𝑆 ⊆ ℂ)
 
Theoremelply 14880* Definition of a polynomial with coefficients in 𝑆. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) ↔ (𝑆 ⊆ ℂ ∧ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑆 ∪ {0}) ↑𝑚0)𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘)))))
 
Theoremelply2 14881* The coefficient function can be assumed to have zeroes outside 0...𝑛. (Contributed by Mario Carneiro, 20-Jul-2014.) (Revised by Mario Carneiro, 23-Aug-2014.)
(𝐹 ∈ (Poly‘𝑆) ↔ (𝑆 ⊆ ℂ ∧ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑆 ∪ {0}) ↑𝑚0)((𝑎 “ (ℤ‘(𝑛 + 1))) = {0} ∧ 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘))))))
 
Theoremplyun0 14882 The set of polynomials is unaffected by the addition of zero. (This is built into the definition because all higher powers of a polynomial are effectively zero, so we require that the coefficient field contain zero to simplify some of our closure theorems.) (Contributed by Mario Carneiro, 17-Jul-2014.)
(Poly‘(𝑆 ∪ {0})) = (Poly‘𝑆)
 
Theoremplyf 14883 A polynomial is a function on the complex numbers. (Contributed by Mario Carneiro, 22-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) → 𝐹:ℂ⟶ℂ)
 
Theoremplyss 14884 The polynomial set function preserves the subset relation. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆𝑇𝑇 ⊆ ℂ) → (Poly‘𝑆) ⊆ (Poly‘𝑇))
 
Theoremplyssc 14885 Every polynomial ring is contained in the ring of polynomials over . (Contributed by Mario Carneiro, 22-Jul-2014.)
(Poly‘𝑆) ⊆ (Poly‘ℂ)
 
Theoremelplyr 14886* 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‘𝑆))
 
Theoremelplyd 14887* Sufficient condition for elementhood in the set of polynomials. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝜑𝑆 ⊆ ℂ)    &   (𝜑𝑁 ∈ ℕ0)    &   ((𝜑𝑘 ∈ (0...𝑁)) → 𝐴𝑆)       (𝜑 → (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(𝐴 · (𝑧𝑘))) ∈ (Poly‘𝑆))
 
Theoremply1termlem 14888* Lemma for ply1term 14889. (Contributed by Mario Carneiro, 26-Jul-2014.)
𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧𝑁)))       ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(if(𝑘 = 𝑁, 𝐴, 0) · (𝑧𝑘))))
 
Theoremply1term 14889* A one-term polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧𝑁)))       ((𝑆 ⊆ ℂ ∧ 𝐴𝑆𝑁 ∈ ℕ0) → 𝐹 ∈ (Poly‘𝑆))
 
Theoremplypow 14890* A power is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆𝑁 ∈ ℕ0) → (𝑧 ∈ ℂ ↦ (𝑧𝑁)) ∈ (Poly‘𝑆))
 
Theoremplyconst 14891 A constant function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 𝐴𝑆) → (ℂ × {𝐴}) ∈ (Poly‘𝑆))
 
Theoremplyid 14892 The identity function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆) → Xp ∈ (Poly‘𝑆))
 
Theoremplyaddlem1 14893* 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(𝑀𝑁, 𝑁, 𝑀))(((𝐴𝑓 + 𝐵)‘𝑘) · (𝑧𝑘))))
 
Theoremplymullem1 14894* 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...𝑛)((𝐴𝑘) · (𝐵‘(𝑛𝑘))) · (𝑧𝑛))))
 
Theoremplyaddlem 14895* Lemma for plyadd 14897. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   (𝜑𝑀 ∈ ℕ0)    &   (𝜑𝑁 ∈ ℕ0)    &   (𝜑𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑 → (𝐴 “ (ℤ‘(𝑀 + 1))) = {0})    &   (𝜑 → (𝐵 “ (ℤ‘(𝑁 + 1))) = {0})    &   (𝜑𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴𝑘) · (𝑧𝑘))))    &   (𝜑𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵𝑘) · (𝑧𝑘))))       (𝜑 → (𝐹𝑓 + 𝐺) ∈ (Poly‘𝑆))
 
Theoremplymullem 14896* Lemma for plymul 14898. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   (𝜑𝑀 ∈ ℕ0)    &   (𝜑𝑁 ∈ ℕ0)    &   (𝜑𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑 → (𝐴 “ (ℤ‘(𝑀 + 1))) = {0})    &   (𝜑 → (𝐵 “ (ℤ‘(𝑁 + 1))) = {0})    &   (𝜑𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴𝑘) · (𝑧𝑘))))    &   (𝜑𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵𝑘) · (𝑧𝑘))))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 · 𝐺) ∈ (Poly‘𝑆))
 
Theoremplyadd 14897* The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 + 𝐺) ∈ (Poly‘𝑆))
 
Theoremplymul 14898* The product of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 · 𝐺) ∈ (Poly‘𝑆))
 
Theoremplysub 14899* The difference of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)    &   (𝜑 → -1 ∈ 𝑆)       (𝜑 → (𝐹𝑓𝐺) ∈ (Poly‘𝑆))
 
Theoremplyaddcl 14900 The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.)
((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹𝑓 + 𝐺) ∈ (Poly‘ℂ))
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10200 103 10201-10300 104 10301-10400 105 10401-10500 106 10501-10600 107 10601-10700 108 10701-10800 109 10801-10900 110 10901-11000 111 11001-11100 112 11101-11200 113 11201-11300 114 11301-11400 115 11401-11500 116 11501-11600 117 11601-11700 118 11701-11800 119 11801-11900 120 11901-12000 121 12001-12100 122 12101-12200 123 12201-12300 124 12301-12400 125 12401-12500 126 12501-12600 127 12601-12700 128 12701-12800 129 12801-12900 130 12901-13000 131 13001-13100 132 13101-13200 133 13201-13300 134 13301-13400 135 13401-13500 136 13501-13600 137 13601-13700 138 13701-13800 139 13801-13900 140 13901-14000 141 14001-14100 142 14101-14200 143 14201-14300 144 14301-14400 145 14401-14500 146 14501-14600 147 14601-14700 148 14701-14800 149 14801-14900 150 14901-15000 151 15001-15100 152 15101-15200 153 15201-15300 154 15301-15400 155 15401-15500 156 15501-15574
  Copyright terms: Public domain < Previous  Next >