Derivatives of the tan
and arctan
functions. #
Continuity and derivatives of the tangent and arctangent functions.
theorem
Real.hasStrictDerivAt_tan
{x : ℝ}
(h : Real.cos x ≠ 0)
:
HasStrictDerivAt Real.tan (1 / Real.cos x ^ 2) x
theorem
Real.hasDerivAt_tan
{x : ℝ}
(h : Real.cos x ≠ 0)
:
HasDerivAt Real.tan (1 / Real.cos x ^ 2) x
theorem
Real.tendsto_abs_tan_of_cos_eq_zero
{x : ℝ}
(hx : Real.cos x = 0)
:
Filter.Tendsto (fun (x : ℝ) => |Real.tan x|) (nhdsWithin x {x}ᶜ) Filter.atTop
@[simp]
Lemmas for derivatives of the composition of Real.arctan
with a differentiable function #
In this section we register lemmas for the derivatives of the composition of Real.arctan
with a
differentiable function, for standalone use and use with simp
.
theorem
HasStrictDerivAt.arctan
{f : ℝ → ℝ}
{f' : ℝ}
{x : ℝ}
(hf : HasStrictDerivAt f f' x)
:
HasStrictDerivAt (fun (x : ℝ) => Real.arctan (f x)) (1 / (1 + f x ^ 2) * f') x
theorem
HasDerivAt.arctan
{f : ℝ → ℝ}
{f' : ℝ}
{x : ℝ}
(hf : HasDerivAt f f' x)
:
HasDerivAt (fun (x : ℝ) => Real.arctan (f x)) (1 / (1 + f x ^ 2) * f') x
theorem
HasDerivWithinAt.arctan
{f : ℝ → ℝ}
{f' : ℝ}
{x : ℝ}
{s : Set ℝ}
(hf : HasDerivWithinAt f f' s x)
:
HasDerivWithinAt (fun (x : ℝ) => Real.arctan (f x)) (1 / (1 + f x ^ 2) * f') s x
theorem
derivWithin_arctan
{f : ℝ → ℝ}
{x : ℝ}
{s : Set ℝ}
(hf : DifferentiableWithinAt ℝ f s x)
(hxs : UniqueDiffWithinAt ℝ s x)
:
derivWithin (fun (x : ℝ) => Real.arctan (f x)) s x = 1 / (1 + f x ^ 2) * derivWithin f s x
theorem
HasStrictFDerivAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{f' : E →L[ℝ] ℝ}
{x : E}
(hf : HasStrictFDerivAt f f' x)
:
HasStrictFDerivAt (fun (x : E) => Real.arctan (f x)) ((1 / (1 + f x ^ 2)) • f') x
theorem
HasFDerivAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{f' : E →L[ℝ] ℝ}
{x : E}
(hf : HasFDerivAt f f' x)
:
HasFDerivAt (fun (x : E) => Real.arctan (f x)) ((1 / (1 + f x ^ 2)) • f') x
theorem
HasFDerivWithinAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{f' : E →L[ℝ] ℝ}
{x : E}
{s : Set E}
(hf : HasFDerivWithinAt f f' s x)
:
HasFDerivWithinAt (fun (x : E) => Real.arctan (f x)) ((1 / (1 + f x ^ 2)) • f') s x
theorem
fderivWithin_arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
{s : Set E}
(hf : DifferentiableWithinAt ℝ f s x)
(hxs : UniqueDiffWithinAt ℝ s x)
:
fderivWithin ℝ (fun (x : E) => Real.arctan (f x)) s x = (1 / (1 + f x ^ 2)) • fderivWithin ℝ f s x
@[simp]
theorem
fderiv_arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
(hc : DifferentiableAt ℝ f x)
:
theorem
DifferentiableWithinAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
{s : Set E}
(hf : DifferentiableWithinAt ℝ f s x)
:
DifferentiableWithinAt ℝ (fun (x : E) => Real.arctan (f x)) s x
@[simp]
theorem
DifferentiableAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
(hc : DifferentiableAt ℝ f x)
:
DifferentiableAt ℝ (fun (x : E) => Real.arctan (f x)) x
theorem
DifferentiableOn.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{s : Set E}
(hc : DifferentiableOn ℝ f s)
:
DifferentiableOn ℝ (fun (x : E) => Real.arctan (f x)) s
@[simp]
theorem
Differentiable.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
(hc : Differentiable ℝ f)
:
Differentiable ℝ fun (x : E) => Real.arctan (f x)
theorem
ContDiffAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
{n : ℕ∞}
(h : ContDiffAt ℝ n f x)
:
ContDiffAt ℝ n (fun (x : E) => Real.arctan (f x)) x
theorem
ContDiff.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{n : ℕ∞}
(h : ContDiff ℝ n f)
:
ContDiff ℝ n fun (x : E) => Real.arctan (f x)
theorem
ContDiffWithinAt.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
{s : Set E}
{n : ℕ∞}
(h : ContDiffWithinAt ℝ n f s x)
:
ContDiffWithinAt ℝ n (fun (x : E) => Real.arctan (f x)) s x
theorem
ContDiffOn.arctan
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{s : Set E}
{n : ℕ∞}
(h : ContDiffOn ℝ n f s)
:
ContDiffOn ℝ n (fun (x : E) => Real.arctan (f x)) s