HomeHome Intuitionistic Logic Explorer
Theorem List (p. 150 of 157)
< 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 - 14901-15000   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremdvmptnegcn 14901* 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 14902* Function-builder for derivative, subtraction rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.)
((𝜑𝑥 ∈ ℂ) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐵𝑉)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵))    &   ((𝜑𝑥 ∈ ℂ) → 𝐶 ∈ ℂ)    &   ((𝜑𝑥 ∈ ℂ) → 𝐷𝑊)    &   (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐶)) = (𝑥 ∈ ℂ ↦ 𝐷))       (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ (𝐴𝐶))) = (𝑥 ∈ ℂ ↦ (𝐵𝐷)))
 
Theoremdvmptcjx 14903* Function-builder for derivative, conjugate rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 24-May-2024.)
((𝜑𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑥𝑋) → 𝐵𝑉)    &   (𝜑 → (ℝ D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))    &   (𝜑𝑋 ⊆ ℝ)       (𝜑 → (ℝ D (𝑥𝑋 ↦ (∗‘𝐴))) = (𝑥𝑋 ↦ (∗‘𝐵)))
 
Theoremdvmptfsum 14904* Function-builder for derivative, finite sums rule. (Contributed by Stefan O'Rear, 12-Nov-2014.)
𝐽 = (𝐾t 𝑆)    &   𝐾 = (TopOpen‘ℂfld)    &   (𝜑𝑆 ∈ {ℝ, ℂ})    &   (𝜑𝑋𝐽)    &   (𝜑𝐼 ∈ Fin)    &   ((𝜑𝑖𝐼𝑥𝑋) → 𝐴 ∈ ℂ)    &   ((𝜑𝑖𝐼𝑥𝑋) → 𝐵 ∈ ℂ)    &   ((𝜑𝑖𝐼) → (𝑆 D (𝑥𝑋𝐴)) = (𝑥𝑋𝐵))       (𝜑 → (𝑆 D (𝑥𝑋 ↦ Σ𝑖𝐼 𝐴)) = (𝑥𝑋 ↦ Σ𝑖𝐼 𝐵))
 
Theoremdveflem 14905 Derivative of the exponential function at 0. The key step in the proof is eftlub 11836, 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 14906 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 14907 Extend class notation to include the set of complex polynomials.
class Poly
 
Syntaxcidp 14908 Extend class notation to include the identity polynomial.
class Xp
 
Definitiondf-ply 14909* 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 14910 Define the identity polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
Xp = ( I ↾ ℂ)
 
Theoremplyval 14911* Value of the polynomial set function. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝑆 ⊆ ℂ → (Poly‘𝑆) = {𝑓 ∣ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑆 ∪ {0}) ↑𝑚0)𝑓 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘)))})
 
Theoremplybss 14912 Reverse closure of the parameter 𝑆 of the polynomial set function. (Contributed by Mario Carneiro, 22-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) → 𝑆 ⊆ ℂ)
 
Theoremelply 14913* Definition of a polynomial with coefficients in 𝑆. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) ↔ (𝑆 ⊆ ℂ ∧ ∃𝑛 ∈ ℕ0𝑎 ∈ ((𝑆 ∪ {0}) ↑𝑚0)𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑛)((𝑎𝑘) · (𝑧𝑘)))))
 
Theoremelply2 14914* 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 14915 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 14916 A polynomial is a function on the complex numbers. (Contributed by Mario Carneiro, 22-Jul-2014.)
(𝐹 ∈ (Poly‘𝑆) → 𝐹:ℂ⟶ℂ)
 
Theoremplyss 14917 The polynomial set function preserves the subset relation. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆𝑇𝑇 ⊆ ℂ) → (Poly‘𝑆) ⊆ (Poly‘𝑇))
 
Theoremplyssc 14918 Every polynomial ring is contained in the ring of polynomials over . (Contributed by Mario Carneiro, 22-Jul-2014.)
(Poly‘𝑆) ⊆ (Poly‘ℂ)
 
Theoremelplyr 14919* 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 14920* Sufficient condition for elementhood in the set of polynomials. (Contributed by Mario Carneiro, 17-Jul-2014.)
(𝜑𝑆 ⊆ ℂ)    &   (𝜑𝑁 ∈ ℕ0)    &   ((𝜑𝑘 ∈ (0...𝑁)) → 𝐴𝑆)       (𝜑 → (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(𝐴 · (𝑧𝑘))) ∈ (Poly‘𝑆))
 
Theoremply1termlem 14921* Lemma for ply1term 14922. (Contributed by Mario Carneiro, 26-Jul-2014.)
𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧𝑁)))       ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → 𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(if(𝑘 = 𝑁, 𝐴, 0) · (𝑧𝑘))))
 
Theoremply1term 14922* A one-term polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
𝐹 = (𝑧 ∈ ℂ ↦ (𝐴 · (𝑧𝑁)))       ((𝑆 ⊆ ℂ ∧ 𝐴𝑆𝑁 ∈ ℕ0) → 𝐹 ∈ (Poly‘𝑆))
 
Theoremplypow 14923* A power is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆𝑁 ∈ ℕ0) → (𝑧 ∈ ℂ ↦ (𝑧𝑁)) ∈ (Poly‘𝑆))
 
Theoremplyconst 14924 A constant function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 𝐴𝑆) → (ℂ × {𝐴}) ∈ (Poly‘𝑆))
 
Theoremplyid 14925 The identity function is a polynomial. (Contributed by Mario Carneiro, 17-Jul-2014.)
((𝑆 ⊆ ℂ ∧ 1 ∈ 𝑆) → Xp ∈ (Poly‘𝑆))
 
Theoremplyaddlem1 14926* 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 14927* 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 14928* Lemma for plyadd 14930. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   (𝜑𝑀 ∈ ℕ0)    &   (𝜑𝑁 ∈ ℕ0)    &   (𝜑𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑 → (𝐴 “ (ℤ‘(𝑀 + 1))) = {0})    &   (𝜑 → (𝐵 “ (ℤ‘(𝑁 + 1))) = {0})    &   (𝜑𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴𝑘) · (𝑧𝑘))))    &   (𝜑𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵𝑘) · (𝑧𝑘))))       (𝜑 → (𝐹𝑓 + 𝐺) ∈ (Poly‘𝑆))
 
Theoremplymullem 14929* Lemma for plymul 14931. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   (𝜑𝑀 ∈ ℕ0)    &   (𝜑𝑁 ∈ ℕ0)    &   (𝜑𝐴 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑𝐵 ∈ ((𝑆 ∪ {0}) ↑𝑚0))    &   (𝜑 → (𝐴 “ (ℤ‘(𝑀 + 1))) = {0})    &   (𝜑 → (𝐵 “ (ℤ‘(𝑁 + 1))) = {0})    &   (𝜑𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑀)((𝐴𝑘) · (𝑧𝑘))))    &   (𝜑𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐵𝑘) · (𝑧𝑘))))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 · 𝐺) ∈ (Poly‘𝑆))
 
Theoremplyadd 14930* The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 + 𝐺) ∈ (Poly‘𝑆))
 
Theoremplymul 14931* The product of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)       (𝜑 → (𝐹𝑓 · 𝐺) ∈ (Poly‘𝑆))
 
Theoremplysub 14932* The difference of two polynomials is a polynomial. (Contributed by Mario Carneiro, 21-Jul-2014.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)    &   (𝜑 → -1 ∈ 𝑆)       (𝜑 → (𝐹𝑓𝐺) ∈ (Poly‘𝑆))
 
Theoremplyaddcl 14933 The sum of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.)
((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹𝑓 + 𝐺) ∈ (Poly‘ℂ))
 
Theoremplymulcl 14934 The product of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.)
((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹𝑓 · 𝐺) ∈ (Poly‘ℂ))
 
Theoremplysubcl 14935 The difference of two polynomials is a polynomial. (Contributed by Mario Carneiro, 24-Jul-2014.)
((𝐹 ∈ (Poly‘𝑆) ∧ 𝐺 ∈ (Poly‘𝑆)) → (𝐹𝑓𝐺) ∈ (Poly‘ℂ))
 
Theoremplycolemc 14936* Lemma for plyco 14937. The result expressed as a sum, with a degree and coefficients for 𝐹 specified as hypotheses. (Contributed by Jim Kingdon, 20-Sep-2025.)
(𝜑𝐹 ∈ (Poly‘𝑆))    &   (𝜑𝐺 ∈ (Poly‘𝑆))    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 + 𝑦) ∈ 𝑆)    &   ((𝜑 ∧ (𝑥𝑆𝑦𝑆)) → (𝑥 · 𝑦) ∈ 𝑆)    &   (𝜑𝑁 ∈ ℕ0)    &   (𝜑𝐴:ℕ0⟶(𝑆 ∪ {0}))    &   (𝜑 → (𝐴 “ (ℤ‘(𝑁 + 1))) = {0})    &   (𝜑𝐹 = (𝑥 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝑥𝑘))))       (𝜑 → (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · ((𝐺𝑧)↑𝑘))) ∈ (Poly‘𝑆))
 
Theoremplyco 14937* 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‘𝑆))
 
Theoremplycjlemc 14938* Lemma for plycj 14939. (Contributed by Mario Carneiro, 24-Jul-2014.) (Revised by Jim Kingdon, 22-Sep-2025.)
(𝜑𝑁 ∈ ℕ0)    &   𝐺 = ((∗ ∘ 𝐹) ∘ ∗)    &   (𝜑𝐴:ℕ0⟶(𝑆 ∪ {0}))    &   (𝜑𝐹 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)((𝐴𝑘) · (𝑧𝑘))))    &   (𝜑𝐹 ∈ (Poly‘𝑆))       (𝜑𝐺 = (𝑧 ∈ ℂ ↦ Σ𝑘 ∈ (0...𝑁)(((∗ ∘ 𝐴)‘𝑘) · (𝑧𝑘))))
 
Theoremplycj 14939* 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.)
𝐺 = ((∗ ∘ 𝐹) ∘ ∗)    &   ((𝜑𝑥𝑆) → (∗‘𝑥) ∈ 𝑆)    &   (𝜑𝐹 ∈ (Poly‘𝑆))       (𝜑𝐺 ∈ (Poly‘𝑆))
 
Theoremplycn 14940 A polynomial is a continuous function. (Contributed by Mario Carneiro, 23-Jul-2014.) Avoid ax-mulf 7997. (Revised by GG, 16-Mar-2025.)
(𝐹 ∈ (Poly‘𝑆) → 𝐹 ∈ (ℂ–cn→ℂ))
 
Theoremplyrecj 14941 A polynomial with real coefficients distributes under conjugation. (Contributed by Mario Carneiro, 24-Jul-2014.)
((𝐹 ∈ (Poly‘ℝ) ∧ 𝐴 ∈ ℂ) → (∗‘(𝐹𝐴)) = (𝐹‘(∗‘𝐴)))
 
Theoremplyreres 14942 Real-coefficient polynomials restrict to real functions. (Contributed by Stefan O'Rear, 16-Nov-2014.)
(𝐹 ∈ (Poly‘ℝ) → (𝐹 ↾ ℝ):ℝ⟶ℝ)
 
Theoremdvply1 14943* 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 𝐹) = 𝐺)
 
11.2  Basic trigonometry
 
11.2.1  The exponential, sine, and cosine functions (cont.)
 
Theoremefcn 14944 The exponential function is continuous. (Contributed by Paul Chapman, 15-Sep-2007.) (Revised by Mario Carneiro, 20-Jun-2015.)
exp ∈ (ℂ–cn→ℂ)
 
Theoremsincn 14945 Sine is continuous. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Mario Carneiro, 3-Sep-2014.)
sin ∈ (ℂ–cn→ℂ)
 
Theoremcoscn 14946 Cosine is continuous. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Mario Carneiro, 3-Sep-2014.)
cos ∈ (ℂ–cn→ℂ)
 
Theoremreeff1olem 14947* Lemma for reeff1o 14949. (Contributed by Paul Chapman, 18-Oct-2007.) (Revised by Mario Carneiro, 30-Apr-2014.)
((𝑈 ∈ ℝ ∧ 1 < 𝑈) → ∃𝑥 ∈ ℝ (exp‘𝑥) = 𝑈)
 
Theoremreeff1oleme 14948* Lemma for reeff1o 14949. (Contributed by Jim Kingdon, 15-May-2024.)
(𝑈 ∈ (0(,)e) → ∃𝑥 ∈ ℝ (exp‘𝑥) = 𝑈)
 
Theoremreeff1o 14949 The real exponential function is one-to-one onto. (Contributed by Paul Chapman, 18-Oct-2007.) (Revised by Mario Carneiro, 10-Nov-2013.)
(exp ↾ ℝ):ℝ–1-1-onto→ℝ+
 
Theoremefltlemlt 14950 Lemma for eflt 14951. The converse of efltim 11844 plus the epsilon-delta setup. (Contributed by Jim Kingdon, 22-May-2024.)
(𝜑𝐴 ∈ ℝ)    &   (𝜑𝐵 ∈ ℝ)    &   (𝜑 → (exp‘𝐴) < (exp‘𝐵))    &   (𝜑𝐷 ∈ ℝ+)    &   (𝜑 → ((abs‘(𝐴𝐵)) < 𝐷 → (abs‘((exp‘𝐴) − (exp‘𝐵))) < ((exp‘𝐵) − (exp‘𝐴))))       (𝜑𝐴 < 𝐵)
 
Theoremeflt 14951 The exponential function on the reals is strictly increasing. (Contributed by Paul Chapman, 21-Aug-2007.) (Revised by Jim Kingdon, 21-May-2024.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 < 𝐵 ↔ (exp‘𝐴) < (exp‘𝐵)))
 
Theoremefle 14952 The exponential function on the reals is nondecreasing. (Contributed by Mario Carneiro, 11-Mar-2014.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴𝐵 ↔ (exp‘𝐴) ≤ (exp‘𝐵)))
 
Theoremreefiso 14953 The exponential function on the reals determines an isomorphism from reals onto positive reals. (Contributed by Steve Rodriguez, 25-Nov-2007.) (Revised by Mario Carneiro, 11-Mar-2014.)
(exp ↾ ℝ) Isom < , < (ℝ, ℝ+)
 
Theoremreapef 14954 Apartness and the exponential function for reals. (Contributed by Jim Kingdon, 11-Jul-2024.)
((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 # 𝐵 ↔ (exp‘𝐴) # (exp‘𝐵)))
 
11.2.2  Properties of pi = 3.14159...
 
Theorempilem1 14955 Lemma for pire , pigt2lt4 and sinpi . (Contributed by Mario Carneiro, 9-May-2014.)
(𝐴 ∈ (ℝ+ ∩ (sin “ {0})) ↔ (𝐴 ∈ ℝ+ ∧ (sin‘𝐴) = 0))
 
Theoremcosz12 14956 Cosine has a zero between 1 and 2. (Contributed by Mario Carneiro and Jim Kingdon, 7-Mar-2024.)
𝑝 ∈ (1(,)2)(cos‘𝑝) = 0
 
Theoremsin0pilem1 14957* Lemma for pi related theorems. (Contributed by Mario Carneiro and Jim Kingdon, 8-Mar-2024.)
𝑝 ∈ (1(,)2)((cos‘𝑝) = 0 ∧ ∀𝑥 ∈ (𝑝(,)(2 · 𝑝))0 < (sin‘𝑥))
 
Theoremsin0pilem2 14958* Lemma for pi related theorems. (Contributed by Mario Carneiro and Jim Kingdon, 8-Mar-2024.)
𝑞 ∈ (2(,)4)((sin‘𝑞) = 0 ∧ ∀𝑥 ∈ (0(,)𝑞)0 < (sin‘𝑥))
 
Theorempilem3 14959 Lemma for pi related theorems. (Contributed by Jim Kingdon, 9-Mar-2024.)
(π ∈ (2(,)4) ∧ (sin‘π) = 0)
 
Theorempigt2lt4 14960 π is between 2 and 4. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 9-May-2014.)
(2 < π ∧ π < 4)
 
Theoremsinpi 14961 The sine of π is 0. (Contributed by Paul Chapman, 23-Jan-2008.)
(sin‘π) = 0
 
Theorempire 14962 π is a real number. (Contributed by Paul Chapman, 23-Jan-2008.)
π ∈ ℝ
 
Theorempicn 14963 π is a complex number. (Contributed by David A. Wheeler, 6-Dec-2018.)
π ∈ ℂ
 
Theorempipos 14964 π is positive. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 9-May-2014.)
0 < π
 
Theorempirp 14965 π is a positive real. (Contributed by Glauco Siliprandi, 11-Dec-2019.)
π ∈ ℝ+
 
Theoremnegpicn 14966 is a real number. (Contributed by David A. Wheeler, 8-Dec-2018.)
-π ∈ ℂ
 
Theoremsinhalfpilem 14967 Lemma for sinhalfpi 14972 and coshalfpi 14973. (Contributed by Paul Chapman, 23-Jan-2008.)
((sin‘(π / 2)) = 1 ∧ (cos‘(π / 2)) = 0)
 
Theoremhalfpire 14968 π / 2 is real. (Contributed by David Moews, 28-Feb-2017.)
(π / 2) ∈ ℝ
 
Theoremneghalfpire 14969 -π / 2 is real. (Contributed by David A. Wheeler, 8-Dec-2018.)
-(π / 2) ∈ ℝ
 
Theoremneghalfpirx 14970 -π / 2 is an extended real. (Contributed by David A. Wheeler, 8-Dec-2018.)
-(π / 2) ∈ ℝ*
 
Theorempidiv2halves 14971 Adding π / 2 to itself gives π. See 2halves 9214. (Contributed by David A. Wheeler, 8-Dec-2018.)
((π / 2) + (π / 2)) = π
 
Theoremsinhalfpi 14972 The sine of π / 2 is 1. (Contributed by Paul Chapman, 23-Jan-2008.)
(sin‘(π / 2)) = 1
 
Theoremcoshalfpi 14973 The cosine of π / 2 is 0. (Contributed by Paul Chapman, 23-Jan-2008.)
(cos‘(π / 2)) = 0
 
Theoremcosneghalfpi 14974 The cosine of -π / 2 is zero. (Contributed by David Moews, 28-Feb-2017.)
(cos‘-(π / 2)) = 0
 
Theoremefhalfpi 14975 The exponential of iπ / 2 is i. (Contributed by Mario Carneiro, 9-May-2014.)
(exp‘(i · (π / 2))) = i
 
Theoremcospi 14976 The cosine of π is -1. (Contributed by Paul Chapman, 23-Jan-2008.)
(cos‘π) = -1
 
Theoremefipi 14977 The exponential of i · π is -1. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
(exp‘(i · π)) = -1
 
Theoremeulerid 14978 Euler's identity. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 9-May-2014.)
((exp‘(i · π)) + 1) = 0
 
Theoremsin2pi 14979 The sine of is 0. (Contributed by Paul Chapman, 23-Jan-2008.)
(sin‘(2 · π)) = 0
 
Theoremcos2pi 14980 The cosine of is 1. (Contributed by Paul Chapman, 23-Jan-2008.)
(cos‘(2 · π)) = 1
 
Theoremef2pi 14981 The exponential of 2πi is 1. (Contributed by Mario Carneiro, 9-May-2014.)
(exp‘(i · (2 · π))) = 1
 
Theoremef2kpi 14982 If 𝐾 is an integer, then the exponential of 2𝐾πi is 1. (Contributed by Mario Carneiro, 9-May-2014.)
(𝐾 ∈ ℤ → (exp‘((i · (2 · π)) · 𝐾)) = 1)
 
Theoremefper 14983 The exponential function is periodic. (Contributed by Paul Chapman, 21-Apr-2008.) (Proof shortened by Mario Carneiro, 10-May-2014.)
((𝐴 ∈ ℂ ∧ 𝐾 ∈ ℤ) → (exp‘(𝐴 + ((i · (2 · π)) · 𝐾))) = (exp‘𝐴))
 
Theoremsinperlem 14984 Lemma for sinper 14985 and cosper 14986. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
(𝐴 ∈ ℂ → (𝐹𝐴) = (((exp‘(i · 𝐴))𝑂(exp‘(-i · 𝐴))) / 𝐷))    &   ((𝐴 + (𝐾 · (2 · π))) ∈ ℂ → (𝐹‘(𝐴 + (𝐾 · (2 · π)))) = (((exp‘(i · (𝐴 + (𝐾 · (2 · π)))))𝑂(exp‘(-i · (𝐴 + (𝐾 · (2 · π)))))) / 𝐷))       ((𝐴 ∈ ℂ ∧ 𝐾 ∈ ℤ) → (𝐹‘(𝐴 + (𝐾 · (2 · π)))) = (𝐹𝐴))
 
Theoremsinper 14985 The sine function is periodic. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
((𝐴 ∈ ℂ ∧ 𝐾 ∈ ℤ) → (sin‘(𝐴 + (𝐾 · (2 · π)))) = (sin‘𝐴))
 
Theoremcosper 14986 The cosine function is periodic. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
((𝐴 ∈ ℂ ∧ 𝐾 ∈ ℤ) → (cos‘(𝐴 + (𝐾 · (2 · π)))) = (cos‘𝐴))
 
Theoremsin2kpi 14987 If 𝐾 is an integer, then the sine of 2𝐾π is 0. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
(𝐾 ∈ ℤ → (sin‘(𝐾 · (2 · π))) = 0)
 
Theoremcos2kpi 14988 If 𝐾 is an integer, then the cosine of 2𝐾π is 1. (Contributed by Paul Chapman, 23-Jan-2008.) (Revised by Mario Carneiro, 10-May-2014.)
(𝐾 ∈ ℤ → (cos‘(𝐾 · (2 · π))) = 1)
 
Theoremsin2pim 14989 Sine of a number subtracted from 2 · π. (Contributed by Paul Chapman, 15-Mar-2008.)
(𝐴 ∈ ℂ → (sin‘((2 · π) − 𝐴)) = -(sin‘𝐴))
 
Theoremcos2pim 14990 Cosine of a number subtracted from 2 · π. (Contributed by Paul Chapman, 15-Mar-2008.)
(𝐴 ∈ ℂ → (cos‘((2 · π) − 𝐴)) = (cos‘𝐴))
 
Theoremsinmpi 14991 Sine of a number less π. (Contributed by Paul Chapman, 15-Mar-2008.)
(𝐴 ∈ ℂ → (sin‘(𝐴 − π)) = -(sin‘𝐴))
 
Theoremcosmpi 14992 Cosine of a number less π. (Contributed by Paul Chapman, 15-Mar-2008.)
(𝐴 ∈ ℂ → (cos‘(𝐴 − π)) = -(cos‘𝐴))
 
Theoremsinppi 14993 Sine of a number plus π. (Contributed by NM, 10-Aug-2008.)
(𝐴 ∈ ℂ → (sin‘(𝐴 + π)) = -(sin‘𝐴))
 
Theoremcosppi 14994 Cosine of a number plus π. (Contributed by NM, 18-Aug-2008.)
(𝐴 ∈ ℂ → (cos‘(𝐴 + π)) = -(cos‘𝐴))
 
Theoremefimpi 14995 The exponential function at i times a real number less π. (Contributed by Paul Chapman, 15-Mar-2008.)
(𝐴 ∈ ℂ → (exp‘(i · (𝐴 − π))) = -(exp‘(i · 𝐴)))
 
Theoremsinhalfpip 14996 The sine of π / 2 plus a number. (Contributed by Paul Chapman, 24-Jan-2008.)
(𝐴 ∈ ℂ → (sin‘((π / 2) + 𝐴)) = (cos‘𝐴))
 
Theoremsinhalfpim 14997 The sine of π / 2 minus a number. (Contributed by Paul Chapman, 24-Jan-2008.)
(𝐴 ∈ ℂ → (sin‘((π / 2) − 𝐴)) = (cos‘𝐴))
 
Theoremcoshalfpip 14998 The cosine of π / 2 plus a number. (Contributed by Paul Chapman, 24-Jan-2008.)
(𝐴 ∈ ℂ → (cos‘((π / 2) + 𝐴)) = -(sin‘𝐴))
 
Theoremcoshalfpim 14999 The cosine of π / 2 minus a number. (Contributed by Paul Chapman, 24-Jan-2008.)
(𝐴 ∈ ℂ → (cos‘((π / 2) − 𝐴)) = (sin‘𝐴))
 
Theoremptolemy 15000 Ptolemy's Theorem. This theorem is named after the Greek astronomer and mathematician Ptolemy (Claudius Ptolemaeus). This particular version is expressed using the sine function. It is proved by expanding all the multiplication of sines to a product of cosines of differences using sinmul 11890, then using algebraic simplification to show that both sides are equal. This formalization is based on the proof in "Trigonometry" by Gelfand and Saul. This is Metamath 100 proof #95. (Contributed by David A. Wheeler, 31-May-2015.)
(((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) ∧ (𝐶 ∈ ℂ ∧ 𝐷 ∈ ℂ) ∧ ((𝐴 + 𝐵) + (𝐶 + 𝐷)) = π) → (((sin‘𝐴) · (sin‘𝐵)) + ((sin‘𝐶) · (sin‘𝐷))) = ((sin‘(𝐵 + 𝐶)) · (sin‘(𝐴 + 𝐶))))
    < 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-15600 157 15601-15644
  Copyright terms: Public domain < Previous  Next >