Skip to content

Commit 04a303a

Browse files
committed
feat: port Analysis.SpecialFunctions.Trigonometric.ArctanDeriv (#4764)
1 parent b90d4a0 commit 04a303a

File tree

2 files changed

+233
-0
lines changed

2 files changed

+233
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -655,6 +655,7 @@ import Mathlib.Analysis.SpecialFunctions.Pow.Real
655655
import Mathlib.Analysis.SpecialFunctions.Sqrt
656656
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Angle
657657
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Arctan
658+
import Mathlib.Analysis.SpecialFunctions.Trigonometric.ArctanDeriv
658659
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic
659660
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Chebyshev
660661
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Complex
Lines changed: 232 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,232 @@
1+
/-
2+
Copyright (c) 2018 Chris Hughes. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Benjamin Davidson
5+
6+
! This file was ported from Lean 3 source module analysis.special_functions.trigonometric.arctan_deriv
7+
! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Arctan
12+
import Mathlib.Analysis.SpecialFunctions.Trigonometric.ComplexDeriv
13+
14+
/-!
15+
# Derivatives of the `tan` and `arctan` functions.
16+
17+
Continuity and derivatives of the tangent and arctangent functions.
18+
-/
19+
20+
21+
noncomputable section
22+
23+
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue #2220
24+
25+
namespace Real
26+
27+
open Set Filter
28+
29+
open scoped Topology Real
30+
31+
theorem hasStrictDerivAt_tan {x : ℝ} (h : cos x ≠ 0) : HasStrictDerivAt tan (1 / cos x ^ 2) x := by
32+
exact_mod_cast (Complex.hasStrictDerivAt_tan (by exact_mod_cast h)).real_of_complex
33+
#align real.has_strict_deriv_at_tan Real.hasStrictDerivAt_tan
34+
35+
theorem hasDerivAt_tan {x : ℝ} (h : cos x ≠ 0) : HasDerivAt tan (1 / cos x ^ 2) x := by
36+
exact_mod_cast (Complex.hasDerivAt_tan (by exact_mod_cast h)).real_of_complex
37+
#align real.has_deriv_at_tan Real.hasDerivAt_tan
38+
39+
theorem tendsto_abs_tan_of_cos_eq_zero {x : ℝ} (hx : cos x = 0) :
40+
Tendsto (fun x => abs (tan x)) (𝓝[≠] x) atTop := by
41+
have hx : Complex.cos x = 0 := by exact_mod_cast hx
42+
simp only [← Complex.abs_ofReal, Complex.ofReal_tan]
43+
refine' (Complex.tendsto_abs_tan_of_cos_eq_zero hx).comp _
44+
refine' Tendsto.inf Complex.continuous_ofReal.continuousAt _
45+
exact tendsto_principal_principal.2 fun y => mt Complex.ofReal_inj.1
46+
#align real.tendsto_abs_tan_of_cos_eq_zero Real.tendsto_abs_tan_of_cos_eq_zero
47+
48+
theorem tendsto_abs_tan_atTop (k : ℤ) :
49+
Tendsto (fun x => abs (tan x)) (𝓝[≠] ((2 * k + 1) * π / 2)) atTop :=
50+
tendsto_abs_tan_of_cos_eq_zero <| cos_eq_zero_iff.2 ⟨k, rfl⟩
51+
#align real.tendsto_abs_tan_at_top Real.tendsto_abs_tan_atTop
52+
53+
theorem continuousAt_tan {x : ℝ} : ContinuousAt tan x ↔ cos x ≠ 0 := by
54+
refine' ⟨fun hc h₀ => _, fun h => (hasDerivAt_tan h).continuousAt⟩
55+
exact not_tendsto_nhds_of_tendsto_atTop (tendsto_abs_tan_of_cos_eq_zero h₀) _
56+
(hc.norm.tendsto.mono_left inf_le_left)
57+
#align real.continuous_at_tan Real.continuousAt_tan
58+
59+
theorem differentiableAt_tan {x : ℝ} : DifferentiableAt ℝ tan x ↔ cos x ≠ 0 :=
60+
fun h => continuousAt_tan.1 h.continuousAt, fun h => (hasDerivAt_tan h).differentiableAt⟩
61+
#align real.differentiable_at_tan Real.differentiableAt_tan
62+
63+
@[simp]
64+
theorem deriv_tan (x : ℝ) : deriv tan x = 1 / cos x ^ 2 :=
65+
if h : cos x = 0 then by
66+
have : ¬DifferentiableAt ℝ tan x := mt differentiableAt_tan.1 (Classical.not_not.2 h)
67+
simp [deriv_zero_of_not_differentiableAt this, h, sq]
68+
else (hasDerivAt_tan h).deriv
69+
#align real.deriv_tan Real.deriv_tan
70+
71+
@[simp]
72+
theorem contDiffAt_tan {n x} : ContDiffAt ℝ n tan x ↔ cos x ≠ 0 :=
73+
fun h => continuousAt_tan.1 h.continuousAt, fun h =>
74+
(Complex.contDiffAt_tan.2 <| by exact_mod_cast h).real_of_complex⟩
75+
#align real.cont_diff_at_tan Real.contDiffAt_tan
76+
77+
theorem hasDerivAt_tan_of_mem_Ioo {x : ℝ} (h : x ∈ Ioo (-(π / 2) : ℝ) (π / 2)) :
78+
HasDerivAt tan (1 / cos x ^ 2) x :=
79+
hasDerivAt_tan (cos_pos_of_mem_Ioo h).ne'
80+
#align real.has_deriv_at_tan_of_mem_Ioo Real.hasDerivAt_tan_of_mem_Ioo
81+
82+
theorem differentiableAt_tan_of_mem_Ioo {x : ℝ} (h : x ∈ Ioo (-(π / 2) : ℝ) (π / 2)) :
83+
DifferentiableAt ℝ tan x :=
84+
(hasDerivAt_tan_of_mem_Ioo h).differentiableAt
85+
#align real.differentiable_at_tan_of_mem_Ioo Real.differentiableAt_tan_of_mem_Ioo
86+
87+
theorem hasStrictDerivAt_arctan (x : ℝ) : HasStrictDerivAt arctan (1 / (1 + x ^ 2)) x := by
88+
have A : cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne'
89+
simpa [cos_sq_arctan] using
90+
tanLocalHomeomorph.hasStrictDerivAt_symm trivial (by simpa) (hasStrictDerivAt_tan A)
91+
#align real.has_strict_deriv_at_arctan Real.hasStrictDerivAt_arctan
92+
93+
theorem hasDerivAt_arctan (x : ℝ) : HasDerivAt arctan (1 / (1 + x ^ 2)) x :=
94+
(hasStrictDerivAt_arctan x).hasDerivAt
95+
#align real.has_deriv_at_arctan Real.hasDerivAt_arctan
96+
97+
theorem differentiableAt_arctan (x : ℝ) : DifferentiableAt ℝ arctan x :=
98+
(hasDerivAt_arctan x).differentiableAt
99+
#align real.differentiable_at_arctan Real.differentiableAt_arctan
100+
101+
theorem differentiable_arctan : Differentiable ℝ arctan :=
102+
differentiableAt_arctan
103+
#align real.differentiable_arctan Real.differentiable_arctan
104+
105+
@[simp]
106+
theorem deriv_arctan : deriv arctan = fun (x : ℝ) => 1 / (1 + x ^ 2) :=
107+
funext fun x => (hasDerivAt_arctan x).deriv
108+
#align real.deriv_arctan Real.deriv_arctan
109+
110+
theorem contDiff_arctan {n : ℕ∞} : ContDiff ℝ n arctan :=
111+
contDiff_iff_contDiffAt.2 fun x =>
112+
have : cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne'
113+
tanLocalHomeomorph.contDiffAt_symm_deriv (by simpa) trivial (hasDerivAt_tan this)
114+
(contDiffAt_tan.2 this)
115+
#align real.cont_diff_arctan Real.contDiff_arctan
116+
117+
end Real
118+
119+
section
120+
121+
/-!
122+
### Lemmas for derivatives of the composition of `Real.arctan` with a differentiable function
123+
124+
In this section we register lemmas for the derivatives of the composition of `Real.arctan` with a
125+
differentiable function, for standalone use and use with `simp`. -/
126+
127+
128+
open Real
129+
130+
section deriv
131+
132+
variable {f : ℝ → ℝ} {f' x : ℝ} {s : Set ℝ}
133+
134+
theorem HasStrictDerivAt.arctan (hf : HasStrictDerivAt f f' x) :
135+
HasStrictDerivAt (fun x => arctan (f x)) (1 / (1 + f x ^ 2) * f') x :=
136+
(Real.hasStrictDerivAt_arctan (f x)).comp x hf
137+
#align has_strict_deriv_at.arctan HasStrictDerivAt.arctan
138+
139+
theorem HasDerivAt.arctan (hf : HasDerivAt f f' x) :
140+
HasDerivAt (fun x => arctan (f x)) (1 / (1 + f x ^ 2) * f') x :=
141+
(Real.hasDerivAt_arctan (f x)).comp x hf
142+
#align has_deriv_at.arctan HasDerivAt.arctan
143+
144+
theorem HasDerivWithinAt.arctan (hf : HasDerivWithinAt f f' s x) :
145+
HasDerivWithinAt (fun x => arctan (f x)) (1 / (1 + f x ^ 2) * f') s x :=
146+
(Real.hasDerivAt_arctan (f x)).comp_hasDerivWithinAt x hf
147+
#align has_deriv_within_at.arctan HasDerivWithinAt.arctan
148+
149+
theorem derivWithin_arctan (hf : DifferentiableWithinAt ℝ f s x) (hxs : UniqueDiffWithinAt ℝ s x) :
150+
derivWithin (fun x => arctan (f x)) s x = 1 / (1 + f x ^ 2) * derivWithin f s x :=
151+
hf.hasDerivWithinAt.arctan.derivWithin hxs
152+
#align deriv_within_arctan derivWithin_arctan
153+
154+
@[simp]
155+
theorem deriv_arctan (hc : DifferentiableAt ℝ f x) :
156+
deriv (fun x => arctan (f x)) x = 1 / (1 + f x ^ 2) * deriv f x :=
157+
hc.hasDerivAt.arctan.deriv
158+
#align deriv_arctan deriv_arctan
159+
160+
end deriv
161+
162+
section fderiv
163+
164+
variable {E : Type _} [NormedAddCommGroup E] [NormedSpace ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} {x : E}
165+
{s : Set E} {n : ℕ∞}
166+
167+
theorem HasStrictFDerivAt.arctan (hf : HasStrictFDerivAt f f' x) :
168+
HasStrictFDerivAt (fun x => arctan (f x)) ((1 / (1 + f x ^ 2)) • f') x :=
169+
(hasStrictDerivAt_arctan (f x)).comp_hasStrictFDerivAt x hf
170+
#align has_strict_fderiv_at.arctan HasStrictFDerivAt.arctan
171+
172+
theorem HasFDerivAt.arctan (hf : HasFDerivAt f f' x) :
173+
HasFDerivAt (fun x => arctan (f x)) ((1 / (1 + f x ^ 2)) • f') x :=
174+
(hasDerivAt_arctan (f x)).comp_hasFDerivAt x hf
175+
#align has_fderiv_at.arctan HasFDerivAt.arctan
176+
177+
theorem HasFDerivWithinAt.arctan (hf : HasFDerivWithinAt f f' s x) :
178+
HasFDerivWithinAt (fun x => arctan (f x)) ((1 / (1 + f x ^ 2)) • f') s x :=
179+
(hasDerivAt_arctan (f x)).comp_hasFDerivWithinAt x hf
180+
#align has_fderiv_within_at.arctan HasFDerivWithinAt.arctan
181+
182+
theorem fderivWithin_arctan (hf : DifferentiableWithinAt ℝ f s x) (hxs : UniqueDiffWithinAt ℝ s x) :
183+
fderivWithin ℝ (fun x => arctan (f x)) s x = (1 / (1 + f x ^ 2)) • fderivWithin ℝ f s x :=
184+
hf.hasFDerivWithinAt.arctan.fderivWithin hxs
185+
#align fderiv_within_arctan fderivWithin_arctan
186+
187+
@[simp]
188+
theorem fderiv_arctan (hc : DifferentiableAt ℝ f x) :
189+
fderiv ℝ (fun x => arctan (f x)) x = (1 / (1 + f x ^ 2)) • fderiv ℝ f x :=
190+
hc.hasFDerivAt.arctan.fderiv
191+
#align fderiv_arctan fderiv_arctan
192+
193+
theorem DifferentiableWithinAt.arctan (hf : DifferentiableWithinAt ℝ f s x) :
194+
DifferentiableWithinAt ℝ (fun x => Real.arctan (f x)) s x :=
195+
hf.hasFDerivWithinAt.arctan.differentiableWithinAt
196+
#align differentiable_within_at.arctan DifferentiableWithinAt.arctan
197+
198+
@[simp]
199+
theorem DifferentiableAt.arctan (hc : DifferentiableAt ℝ f x) :
200+
DifferentiableAt ℝ (fun x => arctan (f x)) x :=
201+
hc.hasFDerivAt.arctan.differentiableAt
202+
#align differentiable_at.arctan DifferentiableAt.arctan
203+
204+
theorem DifferentiableOn.arctan (hc : DifferentiableOn ℝ f s) :
205+
DifferentiableOn ℝ (fun x => arctan (f x)) s := fun x h => (hc x h).arctan
206+
#align differentiable_on.arctan DifferentiableOn.arctan
207+
208+
@[simp]
209+
theorem Differentiable.arctan (hc : Differentiable ℝ f) : Differentiable ℝ fun x => arctan (f x) :=
210+
fun x => (hc x).arctan
211+
#align differentiable.arctan Differentiable.arctan
212+
213+
theorem ContDiffAt.arctan (h : ContDiffAt ℝ n f x) : ContDiffAt ℝ n (fun x => arctan (f x)) x :=
214+
contDiff_arctan.contDiffAt.comp x h
215+
#align cont_diff_at.arctan ContDiffAt.arctan
216+
217+
theorem ContDiff.arctan (h : ContDiff ℝ n f) : ContDiff ℝ n fun x => arctan (f x) :=
218+
contDiff_arctan.comp h
219+
#align cont_diff.arctan ContDiff.arctan
220+
221+
theorem ContDiffWithinAt.arctan (h : ContDiffWithinAt ℝ n f s x) :
222+
ContDiffWithinAt ℝ n (fun x => arctan (f x)) s x :=
223+
contDiff_arctan.comp_contDiffWithinAt h
224+
#align cont_diff_within_at.arctan ContDiffWithinAt.arctan
225+
226+
theorem ContDiffOn.arctan (h : ContDiffOn ℝ n f s) : ContDiffOn ℝ n (fun x => arctan (f x)) s :=
227+
contDiff_arctan.comp_contDiffOn h
228+
#align cont_diff_on.arctan ContDiffOn.arctan
229+
230+
end fderiv
231+
232+
end

0 commit comments

Comments
 (0)