@@ -301,6 +301,18 @@ theorem le_mul_pow_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) (h : 0 <
301
301
rw [inv_pow, ← div_eq_mul_inv]
302
302
exact hCp n
303
303
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
+
304
316
/-- The radius of the sum of two formal series is at least the minimum of their two radii. -/
305
317
theorem min_radius_le_radius_add (p q : FormalMultilinearSeries 𝕜 E F) :
306
318
min p.radius q.radius ≤ (p + q).radius := by
@@ -387,7 +399,7 @@ every point of `s`. -/
387
399
def AnalyticOnNhd (f : E → F) (s : Set E) :=
388
400
∀ x, x ∈ s → AnalyticAt 𝕜 f x
389
401
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
391
403
this is weaker than `AnalyticOnNhd 𝕜 f s`, as `f` is allowed to be arbitrary outside `s`. -/
392
404
def AnalyticOn (f : E → F) (s : Set E) : Prop :=
393
405
∀ x ∈ s, AnalyticWithinAt 𝕜 f s x
@@ -478,6 +490,16 @@ lemma HasFPowerSeriesWithinOnBall.congr {f g : E → F} {p : FormalMultilinearSe
478
490
refine ⟨hy, ?_⟩
479
491
simpa [edist_eq_coe_nnnorm_sub] using h'y
480
492
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
+
481
503
lemma HasFPowerSeriesWithinAt.congr {f g : E → F} {p : FormalMultilinearSeries 𝕜 E F} {s : Set E}
482
504
{x : E} (h : HasFPowerSeriesWithinAt f p s x) (h' : g =ᶠ[𝓝[s] x] f) (h'' : g x = f x) :
483
505
HasFPowerSeriesWithinAt g p s x := by
@@ -581,6 +603,37 @@ lemma HasFPowerSeriesAt.hasFPowerSeriesWithinAt (hf : HasFPowerSeriesAt f p x) :
581
603
rw [← hasFPowerSeriesWithinAt_univ] at hf
582
604
apply hf.mono (subset_univ _)
583
605
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
+
584
637
theorem HasFPowerSeriesWithinOnBall.coeff_zero (hf : HasFPowerSeriesWithinOnBall f pf s x r)
585
638
(v : Fin 0 → E) : pf 0 v = f x := by
586
639
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
697
750
@[deprecated (since := "2024-09-26")]
698
751
alias analyticOn_congr := analyticOnNhd_congr
699
752
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
+
700
758
lemma AnalyticOn.mono {f : E → F} {s t : Set E} (h : AnalyticOn 𝕜 f t)
701
759
(hs : s ⊆ t) : AnalyticOn 𝕜 f s :=
702
760
fun _ m ↦ (h _ (hs m)).mono hs
703
761
704
762
@[deprecated (since := "2024-09-26")]
705
763
alias AnalyticWithinOn.mono := AnalyticOn.mono
706
764
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
+
707
769
/-!
708
770
### Composition with linear maps
709
771
-/
710
772
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
+
711
785
/-- If a function `f` has a power series `p` on a ball and `g` is linear, then `g ∘ f` has the
712
786
power series `g ∘ p` on the same ball. -/
713
787
theorem ContinuousLinearMap.comp_hasFPowerSeriesOnBall (g : F →L[𝕜] G)
714
788
(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⟩
722
800
723
801
/-- If a function `f` is analytic on a set `s` and `g` is linear, then `g ∘ f` is analytic
724
802
on `s`. -/
@@ -729,9 +807,6 @@ theorem ContinuousLinearMap.comp_analyticOnNhd
729
807
rcases h x hx with ⟨p, r, hp⟩
730
808
exact ⟨g.compFormalMultilinearSeries p, r, g.comp_hasFPowerSeriesOnBall hp⟩
731
809
732
- @[deprecated (since := "2024-09-26")]
733
- alias ContinuousLinearMap.comp_analyticOn := ContinuousLinearMap.comp_analyticOnNhd
734
-
735
810
/-!
736
811
### Relation between analytic function and the partial sums of its power series
737
812
-/
@@ -1228,6 +1303,10 @@ protected theorem FormalMultilinearSeries.hasFPowerSeriesOnBall [CompleteSpace F
1228
1303
rw [zero_add]
1229
1304
exact p.hasSum hy }
1230
1305
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
+
1231
1310
theorem HasFPowerSeriesOnBall.sum (h : HasFPowerSeriesOnBall f p x r) {y : E}
1232
1311
(hy : y ∈ EMetric.ball (0 : E) r) : f (x + y) = p.sum y :=
1233
1312
(h.hasSum hy).tsum_eq.symm
0 commit comments