Skip to content

Commit 7da2960

Browse files
sgouezeljcommelin
andcommitted
feat: expand API around analytic functions (#16985)
More basic lemmas for analytic functions (especially versions for `AnalyticOn` of lemmas we already have for `AnalyticOnNhd`). Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent 418c537 commit 7da2960

File tree

2 files changed

+147
-32
lines changed

2 files changed

+147
-32
lines changed

Mathlib/Analysis/Analytic/Basic.lean

Lines changed: 90 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -301,6 +301,18 @@ theorem le_mul_pow_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) (h : 0 <
301301
rw [inv_pow, ← div_eq_mul_inv]
302302
exact hCp n
303303

304+
lemma radius_le_of_le {𝕜' E' F' : Type*}
305+
[NontriviallyNormedField 𝕜'] [NormedAddCommGroup E'] [NormedSpace 𝕜' E']
306+
[NormedAddCommGroup F'] [NormedSpace 𝕜' F']
307+
{p : FormalMultilinearSeries 𝕜 E F} {q : FormalMultilinearSeries 𝕜' E' F'}
308+
(h : ∀ n, ‖p n‖ ≤ ‖q n‖) : q.radius ≤ p.radius := by
309+
apply le_of_forall_nnreal_lt (fun r hr ↦ ?_)
310+
rcases norm_mul_pow_le_of_lt_radius _ hr with ⟨C, -, hC⟩
311+
apply le_radius_of_bound _ C (fun n ↦ ?_)
312+
apply le_trans _ (hC n)
313+
gcongr
314+
exact h n
315+
304316
/-- The radius of the sum of two formal series is at least the minimum of their two radii. -/
305317
theorem min_radius_le_radius_add (p q : FormalMultilinearSeries 𝕜 E F) :
306318
min p.radius q.radius ≤ (p + q).radius := by
@@ -387,7 +399,7 @@ every point of `s`. -/
387399
def AnalyticOnNhd (f : E → F) (s : Set E) :=
388400
∀ x, x ∈ s → AnalyticAt 𝕜 f x
389401

390-
/-- `f` is analytic within `s` if it is analytic within `s` at each point of `t`. Note that
402+
/-- `f` is analytic within `s` if it is analytic within `s` at each point of `s`. Note that
391403
this is weaker than `AnalyticOnNhd 𝕜 f s`, as `f` is allowed to be arbitrary outside `s`. -/
392404
def AnalyticOn (f : E → F) (s : Set E) : Prop :=
393405
∀ x ∈ s, AnalyticWithinAt 𝕜 f s x
@@ -478,6 +490,16 @@ lemma HasFPowerSeriesWithinOnBall.congr {f g : E → F} {p : FormalMultilinearSe
478490
refine ⟨hy, ?_⟩
479491
simpa [edist_eq_coe_nnnorm_sub] using h'y
480492

493+
/-- Variant of `HasFPowerSeriesWithinOnBall.congr` in which one requests equality on `insert x s`
494+
instead of separating `x` and `s`. -/
495+
lemma HasFPowerSeriesWithinOnBall.congr' {f g : E → F} {p : FormalMultilinearSeries 𝕜 E F}
496+
{s : Set E} {x : E} {r : ℝ≥0∞} (h : HasFPowerSeriesWithinOnBall f p s x r)
497+
(h' : EqOn g f (insert x s ∩ EMetric.ball x r)) :
498+
HasFPowerSeriesWithinOnBall g p s x r := by
499+
refine ⟨h.r_le, h.r_pos, fun {y} hy h'y ↦ ?_⟩
500+
convert h.hasSum hy h'y using 1
501+
exact h' ⟨hy, by simpa [edist_eq_coe_nnnorm_sub] using h'y⟩
502+
481503
lemma HasFPowerSeriesWithinAt.congr {f g : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E}
482504
{x : E} (h : HasFPowerSeriesWithinAt f p s x) (h' : g =ᶠ[𝓝[s] x] f) (h'' : g x = f x) :
483505
HasFPowerSeriesWithinAt g p s x := by
@@ -581,6 +603,37 @@ lemma HasFPowerSeriesAt.hasFPowerSeriesWithinAt (hf : HasFPowerSeriesAt f p x) :
581603
rw [← hasFPowerSeriesWithinAt_univ] at hf
582604
apply hf.mono (subset_univ _)
583605

606+
theorem HasFPowerSeriesWithinAt.mono_of_mem
607+
(h : HasFPowerSeriesWithinAt f p s x) (hst : s ∈ 𝓝[t] x) :
608+
HasFPowerSeriesWithinAt f p t x := by
609+
rcases h with ⟨r, hr⟩
610+
rcases EMetric.mem_nhdsWithin_iff.1 hst with ⟨r', r'_pos, hr'⟩
611+
refine ⟨min r r', ?_⟩
612+
have Z := hr.of_le (by simp [r'_pos, hr.r_pos]) (min_le_left r r')
613+
refine ⟨Z.r_le, Z.r_pos, fun {y} hy h'y ↦ ?_⟩
614+
apply Z.hasSum ?_ h'y
615+
simp only [mem_insert_iff, add_right_eq_self] at hy
616+
rcases hy with rfl | hy
617+
· simp
618+
apply mem_insert_of_mem _ (hr' ?_)
619+
simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, sub_zero, lt_min_iff, mem_inter_iff,
620+
add_sub_cancel_left, hy, and_true] at h'y ⊢
621+
exact h'y.2
622+
623+
@[simp] lemma hasFPowerSeriesWithinOnBall_insert_self :
624+
HasFPowerSeriesWithinOnBall f p (insert x s) x r ↔ HasFPowerSeriesWithinOnBall f p s x r := by
625+
refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ <;>
626+
exact ⟨h.r_le, h.r_pos, fun {y} ↦ by simpa only [insert_idem] using h.hasSum (y := y)⟩
627+
628+
@[simp] theorem hasFPowerSeriesWithinAt_insert {y : E} :
629+
HasFPowerSeriesWithinAt f p (insert y s) x ↔ HasFPowerSeriesWithinAt f p s x := by
630+
rcases eq_or_ne x y with rfl | hy
631+
· simp [HasFPowerSeriesWithinAt]
632+
· refine ⟨fun h ↦ h.mono (subset_insert _ _), fun h ↦ ?_⟩
633+
apply HasFPowerSeriesWithinAt.mono_of_mem h
634+
rw [nhdsWithin_insert_of_ne hy]
635+
exact self_mem_nhdsWithin
636+
584637
theorem HasFPowerSeriesWithinOnBall.coeff_zero (hf : HasFPowerSeriesWithinOnBall f pf s x r)
585638
(v : Fin 0 → E) : pf 0 v = f x := by
586639
have v_eq : v = fun i => 0 := Subsingleton.elim _ _
@@ -697,28 +750,53 @@ theorem analyticOnNhd_congr (hs : IsOpen s) (h : s.EqOn f g) : AnalyticOnNhd
697750
@[deprecated (since := "2024-09-26")]
698751
alias analyticOn_congr := analyticOnNhd_congr
699752

753+
theorem AnalyticWithinAt.mono_of_mem
754+
(h : AnalyticWithinAt 𝕜 f s x) (hst : s ∈ 𝓝[t] x) : AnalyticWithinAt 𝕜 f t x := by
755+
rcases h with ⟨p, hp⟩
756+
exact ⟨p, hp.mono_of_mem hst⟩
757+
700758
lemma AnalyticOn.mono {f : E → F} {s t : Set E} (h : AnalyticOn 𝕜 f t)
701759
(hs : s ⊆ t) : AnalyticOn 𝕜 f s :=
702760
fun _ m ↦ (h _ (hs m)).mono hs
703761

704762
@[deprecated (since := "2024-09-26")]
705763
alias AnalyticWithinOn.mono := AnalyticOn.mono
706764

765+
@[simp] theorem analyticWithinAt_insert {f : E → F} {s : Set E} {x y : E} :
766+
AnalyticWithinAt 𝕜 f (insert y s) x ↔ AnalyticWithinAt 𝕜 f s x := by
767+
simp [AnalyticWithinAt]
768+
707769
/-!
708770
### Composition with linear maps
709771
-/
710772

773+
/-- If a function `f` has a power series `p` on a ball within a set and `g` is linear,
774+
then `g ∘ f` has the power series `g ∘ p` on the same ball. -/
775+
theorem ContinuousLinearMap.comp_hasFPowerSeriesWithinOnBall (g : F →L[𝕜] G)
776+
(h : HasFPowerSeriesWithinOnBall f p s x r) :
777+
HasFPowerSeriesWithinOnBall (g ∘ f) (g.compFormalMultilinearSeries p) s x r where
778+
r_le := h.r_le.trans (p.radius_le_radius_continuousLinearMap_comp _)
779+
r_pos := h.r_pos
780+
hasSum hy h'y := by
781+
simpa only [ContinuousLinearMap.compFormalMultilinearSeries_apply,
782+
ContinuousLinearMap.compContinuousMultilinearMap_coe, Function.comp_apply] using
783+
g.hasSum (h.hasSum hy h'y)
784+
711785
/-- If a function `f` has a power series `p` on a ball and `g` is linear, then `g ∘ f` has the
712786
power series `g ∘ p` on the same ball. -/
713787
theorem ContinuousLinearMap.comp_hasFPowerSeriesOnBall (g : F →L[𝕜] G)
714788
(h : HasFPowerSeriesOnBall f p x r) :
715-
HasFPowerSeriesOnBall (g ∘ f) (g.compFormalMultilinearSeries p) x r :=
716-
{ r_le := h.r_le.trans (p.radius_le_radius_continuousLinearMap_comp _)
717-
r_pos := h.r_pos
718-
hasSum := fun hy => by
719-
simpa only [ContinuousLinearMap.compFormalMultilinearSeries_apply,
720-
ContinuousLinearMap.compContinuousMultilinearMap_coe, Function.comp_apply] using
721-
g.hasSum (h.hasSum hy) }
789+
HasFPowerSeriesOnBall (g ∘ f) (g.compFormalMultilinearSeries p) x r := by
790+
rw [← hasFPowerSeriesWithinOnBall_univ] at h ⊢
791+
exact g.comp_hasFPowerSeriesWithinOnBall h
792+
793+
/-- If a function `f` is analytic on a set `s` and `g` is linear, then `g ∘ f` is analytic
794+
on `s`. -/
795+
theorem ContinuousLinearMap.comp_analyticOn (g : F →L[𝕜] G) (h : AnalyticOn 𝕜 f s) :
796+
AnalyticOn 𝕜 (g ∘ f) s := by
797+
rintro x hx
798+
rcases h x hx with ⟨p, r, hp⟩
799+
exact ⟨g.compFormalMultilinearSeries p, r, g.comp_hasFPowerSeriesWithinOnBall hp⟩
722800

723801
/-- If a function `f` is analytic on a set `s` and `g` is linear, then `g ∘ f` is analytic
724802
on `s`. -/
@@ -729,9 +807,6 @@ theorem ContinuousLinearMap.comp_analyticOnNhd
729807
rcases h x hx with ⟨p, r, hp⟩
730808
exact ⟨g.compFormalMultilinearSeries p, r, g.comp_hasFPowerSeriesOnBall hp⟩
731809

732-
@[deprecated (since := "2024-09-26")]
733-
alias ContinuousLinearMap.comp_analyticOn := ContinuousLinearMap.comp_analyticOnNhd
734-
735810
/-!
736811
### Relation between analytic function and the partial sums of its power series
737812
-/
@@ -1228,6 +1303,10 @@ protected theorem FormalMultilinearSeries.hasFPowerSeriesOnBall [CompleteSpace F
12281303
rw [zero_add]
12291304
exact p.hasSum hy }
12301305

1306+
theorem HasFPowerSeriesWithinOnBall.sum (h : HasFPowerSeriesWithinOnBall f p s x r) {y : E}
1307+
(h'y : x + y ∈ insert x s) (hy : y ∈ EMetric.ball (0 : E) r) : f (x + y) = p.sum y :=
1308+
(h.hasSum h'y hy).tsum_eq.symm
1309+
12311310
theorem HasFPowerSeriesOnBall.sum (h : HasFPowerSeriesOnBall f p x r) {y : E}
12321311
(hy : y ∈ EMetric.ball (0 : E) r) : f (x + y) = p.sum y :=
12331312
(h.hasSum hy).tsum_eq.symm

Mathlib/Analysis/Analytic/ChangeOrigin.lean

Lines changed: 57 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -226,6 +226,12 @@ def derivSeries : FormalMultilinearSeries 𝕜 E (E →L[𝕜] F) :=
226226
(continuousMultilinearCurryFin1 𝕜 E F : (E[×1]→L[𝕜] F) →L[𝕜] E →L[𝕜] F)
227227
|>.compFormalMultilinearSeries (p.changeOriginSeries 1)
228228

229+
theorem radius_le_radius_derivSeries : p.radius ≤ p.derivSeries.radius := by
230+
apply (p.le_changeOriginSeries_radius 1).trans (radius_le_of_le (fun n ↦ ?_))
231+
apply (ContinuousLinearMap.norm_compContinuousMultilinearMap_le _ _).trans
232+
apply mul_le_of_le_one_left (norm_nonneg _)
233+
exact ContinuousLinearMap.opNorm_le_bound _ zero_le_one (by simp)
234+
229235
end
230236

231237
-- From this point on, assume that the space is complete, to make sure that series that converge
@@ -284,39 +290,69 @@ theorem analyticAt_changeOrigin (p : FormalMultilinearSeries 𝕜 E F) (rp : p.r
284290

285291
end FormalMultilinearSeries
286292

293+
287294
section
288295

289-
variable [CompleteSpace F] {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} {x y : E} {r : ℝ≥0∞}
296+
variable [CompleteSpace F] {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E}
297+
{x y : E} {r : ℝ≥0∞}
298+
299+
/-- If a function admits a power series expansion `p` within a set `s` on a ball `B (x, r)`, then
300+
it also admits a power series on any subball of this ball (even with a different center provided
301+
it belongs to `s`), given by `p.changeOrigin`. -/
302+
theorem HasFPowerSeriesWithinOnBall.changeOrigin (hf : HasFPowerSeriesWithinOnBall f p s x r)
303+
(h : (‖y‖₊ : ℝ≥0∞) < r) (hy : x + y ∈ insert x s) :
304+
HasFPowerSeriesWithinOnBall f (p.changeOrigin y) s (x + y) (r - ‖y‖₊) where
305+
r_le := by
306+
apply le_trans _ p.changeOrigin_radius
307+
exact tsub_le_tsub hf.r_le le_rfl
308+
r_pos := by simp [h]
309+
hasSum {z} h'z hz := by
310+
have : f (x + y + z) =
311+
FormalMultilinearSeries.sum (FormalMultilinearSeries.changeOrigin p y) z := by
312+
rw [mem_emetric_ball_zero_iff, lt_tsub_iff_right, add_comm] at hz
313+
rw [p.changeOrigin_eval (hz.trans_le hf.r_le), add_assoc, hf.sum]
314+
· have : insert (x + y) s ⊆ insert (x + y) (insert x s) := by
315+
apply insert_subset_insert (subset_insert _ _)
316+
rw [insert_eq_of_mem hy] at this
317+
apply this
318+
simpa [add_assoc] using h'z
319+
refine mem_emetric_ball_zero_iff.2 (lt_of_le_of_lt ?_ hz)
320+
exact mod_cast nnnorm_add_le y z
321+
rw [this]
322+
apply (p.changeOrigin y).hasSum
323+
refine EMetric.ball_subset_ball (le_trans ?_ p.changeOrigin_radius) hz
324+
exact tsub_le_tsub hf.r_le le_rfl
290325

291326
/-- If a function admits a power series expansion `p` on a ball `B (x, r)`, then it also admits a
292327
power series on any subball of this ball (even with a different center), given by `p.changeOrigin`.
293328
-/
294329
theorem HasFPowerSeriesOnBall.changeOrigin (hf : HasFPowerSeriesOnBall f p x r)
295-
(h : (‖y‖₊ : ℝ≥0∞) < r) : HasFPowerSeriesOnBall f (p.changeOrigin y) (x + y) (r - ‖y‖₊) :=
296-
{ r_le := by
297-
apply le_trans _ p.changeOrigin_radius
298-
exact tsub_le_tsub hf.r_le le_rfl
299-
r_pos := by simp [h]
300-
hasSum := fun {z} hz => by
301-
have : f (x + y + z) =
302-
FormalMultilinearSeries.sum (FormalMultilinearSeries.changeOrigin p y) z := by
303-
rw [mem_emetric_ball_zero_iff, lt_tsub_iff_right, add_comm] at hz
304-
rw [p.changeOrigin_eval (hz.trans_le hf.r_le), add_assoc, hf.sum]
305-
refine mem_emetric_ball_zero_iff.2 (lt_of_le_of_lt ?_ hz)
306-
exact mod_cast nnnorm_add_le y z
307-
rw [this]
308-
apply (p.changeOrigin y).hasSum
309-
refine EMetric.ball_subset_ball (le_trans ?_ p.changeOrigin_radius) hz
310-
exact tsub_le_tsub hf.r_le le_rfl }
330+
(h : (‖y‖₊ : ℝ≥0∞) < r) : HasFPowerSeriesOnBall f (p.changeOrigin y) (x + y) (r - ‖y‖₊) := by
331+
rw [← hasFPowerSeriesWithinOnBall_univ] at hf ⊢
332+
exact hf.changeOrigin h (by simp)
333+
334+
/-- If a function admits a power series expansion `p` on an open ball `B (x, r)`, then
335+
it is analytic at every point of this ball. -/
336+
theorem HasFPowerSeriesWithinOnBall.analyticWithinAt_of_mem
337+
(hf : HasFPowerSeriesWithinOnBall f p s x r)
338+
(h : y ∈ insert x s ∩ EMetric.ball x r) : AnalyticWithinAt 𝕜 f s y := by
339+
have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_coe_nnnorm_sub] using h.2
340+
have := hf.changeOrigin this (by simpa using h.1)
341+
rw [add_sub_cancel] at this
342+
exact this.analyticWithinAt
311343

312344
/-- If a function admits a power series expansion `p` on an open ball `B (x, r)`, then
313345
it is analytic at every point of this ball. -/
314346
theorem HasFPowerSeriesOnBall.analyticAt_of_mem (hf : HasFPowerSeriesOnBall f p x r)
315347
(h : y ∈ EMetric.ball x r) : AnalyticAt 𝕜 f y := by
316-
have : (‖y - x‖₊ : ℝ≥0∞) < r := by simpa [edist_eq_coe_nnnorm_sub] using h
317-
have := hf.changeOrigin this
318-
rw [add_sub_cancel] at this
319-
exact this.analyticAt
348+
rw [← hasFPowerSeriesWithinOnBall_univ] at hf
349+
rw [← analyticWithinAt_univ]
350+
exact hf.analyticWithinAt_of_mem (by simpa using h)
351+
352+
theorem HasFPowerSeriesWithinOnBall.analyticOn (hf : HasFPowerSeriesWithinOnBall f p s x r) :
353+
AnalyticOn 𝕜 f (insert x s ∩ EMetric.ball x r) :=
354+
fun _ hy ↦ ((analyticWithinAt_insert (y := x)).2 (hf.analyticWithinAt_of_mem hy)).mono
355+
inter_subset_left
320356

321357
theorem HasFPowerSeriesOnBall.analyticOnNhd (hf : HasFPowerSeriesOnBall f p x r) :
322358
AnalyticOnNhd 𝕜 f (EMetric.ball x r) :=

0 commit comments

Comments
 (0)