Update analyticAt.lean

This commit is contained in:
Stefan Kebekus 2024-09-12 06:58:43 +02:00
parent dbeb631178
commit f83f772506
1 changed files with 19 additions and 7 deletions

View File

@ -165,26 +165,38 @@ theorem AnalyticAt.order_comp_CLE
obtain ⟨g, h₁g, h₂g, h₃g⟩ := hn obtain ⟨g, h₁g, h₂g, h₃g⟩ := hn
have A := eventually_nhds_comp_composition h₃g .continuous have A := eventually_nhds_comp_composition h₃g .continuous
have t₁ : AnalyticAt (fun z => z - z₀) z₀ := by
apply AnalyticAt.sub
exact ContinuousLinearEquiv.analyticAt z₀
exact analyticAt_const
have t₀ : AnalyticAt (fun z => ( z - z₀) ^ n) z₀ := by
exact pow t₁ n
have : AnalyticAt (fun z ↦ ( z - z₀) ^ n • g ( z) : ) z₀ := by have : AnalyticAt (fun z ↦ ( z - z₀) ^ n • g ( z) : ) z₀ := by
sorry apply AnalyticAt.mul
exact t₀
apply AnalyticAt.comp h₁g
exact ContinuousLinearEquiv.analyticAt z₀
rw [AnalyticAt.order_congr (hf.comp (.analyticAt z₀)) this A] rw [AnalyticAt.order_congr (hf.comp (.analyticAt z₀)) this A]
simp simp
have t₀ : AnalyticAt (fun z => ( z - z₀) ^ n) z₀ := by
sorry
rw [AnalyticAt.order_mul t₀ ((h₁g.comp (.analyticAt z₀)))] rw [AnalyticAt.order_mul t₀ ((h₁g.comp (.analyticAt z₀)))]
have t₁ : AnalyticAt (fun z => z - z₀) z₀ := by
sorry
have : t₁.order = (1 : ) := by have : t₁.order = (1 : ) := by
rw [AnalyticAt.order_eq_nat_iff] rw [AnalyticAt.order_eq_nat_iff]
use (fun z 1) use (fun _ 1)
simp simp
constructor constructor
· exact analyticAt_const · exact analyticAt_const
· apply Filter.Eventually.of_forall · apply Filter.Eventually.of_forall
intro x intro x
sorry calc x - z₀
_ = (x - z₀) := by
exact Eq.symm (ContinuousLinearEquiv.map_sub x z₀)
_ = ((x - z₀) * 1) := by
simp
_ = (x - z₀) * 1 := by
rw [← smul_eq_mul, ← smul_eq_mul]
exact ContinuousLinearEquiv.map_smul (x - z₀) 1
have : t₀.order = n := by have : t₀.order = n := by
rw [AnalyticAt.order_pow t₁, this] rw [AnalyticAt.order_pow t₁, this]