@@ -7,7 +7,7 @@ import Mathlib.Analysis.Analytic.CPolynomial
7
7
import Mathlib.Analysis.Analytic.Inverse
8
8
import Mathlib.Analysis.Analytic.Within
9
9
import Mathlib.Analysis.Calculus.Deriv.Basic
10
- import Mathlib.Analysis.Calculus.ContDiff.Defs
10
+ import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries
11
11
import Mathlib.Analysis.Calculus.FDeriv.Add
12
12
import Mathlib.Analysis.Calculus.FDeriv.Prod
13
13
import Mathlib.Analysis.Normed.Module.Completion
@@ -63,7 +63,7 @@ differentiability at points in a neighborhood of `s`. Therefore, the theorem tha
63
63
64
64
open Filter Asymptotics Set
65
65
66
- open scoped ENNReal Topology ContDiff
66
+ open scoped ENNReal Topology
67
67
68
68
universe u v
69
69
@@ -192,6 +192,7 @@ theorem HasFPowerSeriesOnBall.fderiv_eq [CompleteSpace F] (h : HasFPowerSeriesOn
192
192
fderiv 𝕜 f (x + y) = continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1 ) :=
193
193
(h.hasFDerivAt hy).fderiv
194
194
195
+ /-- If a function has a power series on a ball, then so does its derivative. -/
195
196
protected theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F]
196
197
(h : HasFPowerSeriesOnBall f p x r) :
197
198
HasFPowerSeriesOnBall (fderiv 𝕜 f) p.derivSeries x r := by
@@ -238,7 +239,7 @@ protected theorem AnalyticOnNhd.fderiv [CompleteSpace F] (h : AnalyticOnNhd 𝕜
238
239
alias AnalyticOn.fderiv := AnalyticOnNhd.fderiv
239
240
240
241
/-- If a function is analytic on a set `s`, so are its successive Fréchet derivative. See also
241
- `AnalyticOnNhd.iteratedFDeruv_of_isOpen `, removing the completeness assumption but requiring the set
242
+ `AnalyticOnNhd.iteratedFDeriv_of_isOpen `, removing the completeness assumption but requiring the set
242
243
to be open.-/
243
244
protected theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (n : ℕ) :
244
245
AnalyticOnNhd 𝕜 (iteratedFDeriv 𝕜 n f) s := by
@@ -269,44 +270,6 @@ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F]
269
270
· apply (DifferentiableAt.continuousAt (𝕜 := 𝕜) ?_).continuousWithinAt
270
271
exact (h.iteratedFDeriv m x hx).differentiableAt
271
272
272
- /-- An analytic function is infinitely differentiable. -/
273
- protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s)
274
- {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by
275
- suffices ContDiffOn 𝕜 ω f s from this.of_le le_top
276
- rw [← contDiffOn_infty_iff_contDiffOn_omega]
277
- let t := { x | AnalyticAt 𝕜 f x }
278
- suffices ContDiffOn 𝕜 ∞ f t from this.mono h
279
- have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx
280
- have t_open : IsOpen t := isOpen_analyticAt 𝕜 f
281
- exact contDiffOn_of_continuousOn_differentiableOn
282
- (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr
283
- fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx)
284
- (fun m _ ↦ (H.iteratedFDeriv m).differentiableOn.congr
285
- fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx)
286
-
287
- /-- An analytic function on the whole space is infinitely differentiable there. -/
288
- theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : WithTop ℕ∞} :
289
- ContDiff 𝕜 n f := by
290
- rw [← contDiffOn_univ]
291
- exact h.contDiffOn
292
-
293
- theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : WithTop ℕ∞} :
294
- ContDiffAt 𝕜 n f x := by
295
- obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd
296
- exact hf.contDiffOn.contDiffAt hs
297
-
298
- protected lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E}
299
- (h : AnalyticWithinAt 𝕜 f s x) {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x := by
300
- rcases h.exists_analyticAt with ⟨g, fx, fg, hg⟩
301
- exact hg.contDiffAt.contDiffWithinAt.congr (fg.mono (subset_insert _ _)) fx
302
-
303
- protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E}
304
- (h : AnalyticOn 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s :=
305
- fun x m ↦ (h x m).contDiffWithinAt
306
-
307
- @[deprecated (since := "2024-09-26")]
308
- alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn
309
-
310
273
lemma AnalyticWithinAt.exists_hasFTaylorSeriesUpToOn [CompleteSpace F]
311
274
(n : WithTop ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) :
312
275
∃ u ∈ 𝓝[insert x s] x, ∃ (p : E → FormalMultilinearSeries 𝕜 E F),
@@ -384,7 +347,7 @@ protected theorem AnalyticOn.iteratedFDerivWithin (h : AnalyticOn 𝕜 f s)
384
347
apply AnalyticOnNhd.comp_analyticOn _ (IH.fderivWithin hu) (mapsTo_univ _ _)
385
348
apply LinearIsometryEquiv.analyticOnNhd
386
349
387
- lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞}
350
+ protected lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞}
388
351
(h : AnalyticOn 𝕜 f s) (hu : UniqueDiffOn 𝕜 s) :
389
352
HasFTaylorSeriesUpToOn n f (ftaylorSeriesWithin 𝕜 f s) s := by
390
353
refine ⟨fun x _hx ↦ rfl, fun m _hm x hx ↦ ?_, fun m _hm x hx ↦ ?_⟩
@@ -546,26 +509,6 @@ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) :
546
509
case g => exact ↑(continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (n + 1 ) ↦ E) F).symm
547
510
simp
548
511
549
- /-- A polynomial function is infinitely differentiable. -/
550
- theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} :
551
- ContDiffOn 𝕜 n f s := by
552
- suffices ContDiffOn 𝕜 ω f s from this.of_le le_top
553
- let t := { x | CPolynomialAt 𝕜 f x }
554
- suffices ContDiffOn 𝕜 ω f t from this.mono h
555
- rw [← contDiffOn_infty_iff_contDiffOn_omega]
556
- have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx
557
- have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f
558
- exact contDiffOn_of_continuousOn_differentiableOn
559
- (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr
560
- fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx)
561
- (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr
562
- fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx)
563
-
564
- theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} :
565
- ContDiffAt 𝕜 n f x :=
566
- let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn
567
- hf.contDiffOn.contDiffAt hs
568
-
569
512
end fderiv
570
513
571
514
section deriv
@@ -750,10 +693,6 @@ lemma cPolynomialAt : CPolynomialAt 𝕜 f x :=
750
693
751
694
lemma cPolyomialOn : CPolynomialOn 𝕜 f ⊤ := fun x _ ↦ f.cPolynomialAt x
752
695
753
- lemma contDiffAt : ContDiffAt 𝕜 n f x := (f.cPolynomialAt x).contDiffAt
754
-
755
- lemma contDiff : ContDiff 𝕜 n f := contDiff_iff_contDiffAt.mpr f.contDiffAt
756
-
757
696
end ContinuousMultilinearMap
758
697
759
698
namespace FormalMultilinearSeries
@@ -797,8 +736,8 @@ private theorem factorial_smul' {n : ℕ} : ∀ {F : Type max u v} [NormedAddCom
797
736
n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by
798
737
induction n with | zero => _ | succ n ih => _ <;> intro F _ _ _ p f h
799
738
· rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag]
800
- · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply,
801
- ih h.fderiv, iteratedFDeriv_succ_apply_right]
739
+ · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag,
740
+ ← ContinuousLinearMap.smul_apply, ih h.fderiv, iteratedFDeriv_succ_apply_right]
802
741
rfl
803
742
804
743
variable [CompleteSpace F]
@@ -808,8 +747,8 @@ theorem factorial_smul (n : ℕ) :
808
747
n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by
809
748
cases n
810
749
· rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag]
811
- · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply,
812
- factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right]
750
+ · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag,
751
+ ← ContinuousLinearMap.smul_apply, factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right]
813
752
rfl
814
753
815
754
theorem hasSum_iteratedFDeriv [CharZero 𝕜] {y : E} (hy : y ∈ EMetric.ball 0 r) :
0 commit comments