nevanlinna/Nevanlinna/analyticAt.lean

183 lines
5.6 KiB
Plaintext
Raw Normal View History

2024-09-10 10:45:53 +02:00
import Mathlib.Analysis.Analytic.IsolatedZeros
import Mathlib.Analysis.Complex.Basic
2024-09-11 15:48:43 +02:00
import Mathlib.Analysis.Analytic.Linear
2024-09-10 10:45:53 +02:00
theorem AnalyticAt.order_mul
{f₁ f₂ : }
{z₀ : }
(hf₁ : AnalyticAt f₁ z₀)
(hf₂ : AnalyticAt f₂ z₀) :
(AnalyticAt.mul hf₁ hf₂).order = hf₁.order + hf₂.order := by
by_cases h₂f₁ : hf₁.order =
· simp [h₂f₁]
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff]
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at h₂f₁
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₂f₁
use t
constructor
· intro y hy
rw [h₁t y hy]
ring
· exact ⟨h₂t, h₃t⟩
· by_cases h₂f₂ : hf₂.order =
· simp [h₂f₂]
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff]
rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at h₂f₂
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₂f₂
use t
constructor
· intro y hy
rw [h₁t y hy]
ring
· exact ⟨h₂t, h₃t⟩
· obtain ⟨g₁, h₁g₁, h₂g₁, h₃g₁⟩ := (AnalyticAt.order_eq_nat_iff hf₁ ↑hf₁.order.toNat).1 (eq_comm.1 (ENat.coe_toNat h₂f₁))
obtain ⟨g₂, h₁g₂, h₂g₂, h₃g₂⟩ := (AnalyticAt.order_eq_nat_iff hf₂ ↑hf₂.order.toNat).1 (eq_comm.1 (ENat.coe_toNat h₂f₂))
rw [← ENat.coe_toNat h₂f₁, ← ENat.coe_toNat h₂f₂, ← ENat.coe_add]
rw [AnalyticAt.order_eq_nat_iff (AnalyticAt.mul hf₁ hf₂) ↑(hf₁.order.toNat + hf₂.order.toNat)]
use g₁ * g₂
constructor
· exact AnalyticAt.mul h₁g₁ h₁g₂
· constructor
· simp; tauto
· obtain ⟨t₁, h₁t₁, h₂t₁, h₃t₁⟩ := eventually_nhds_iff.1 h₃g₁
obtain ⟨t₂, h₁t₂, h₂t₂, h₃t₂⟩ := eventually_nhds_iff.1 h₃g₂
rw [eventually_nhds_iff]
use t₁ ∩ t₂
constructor
· intro y hy
rw [h₁t₁ y hy.1, h₁t₂ y hy.2]
simp; ring
· constructor
· exact IsOpen.inter h₂t₁ h₂t₂
· exact Set.mem_inter h₃t₁ h₃t₂
2024-09-10 14:21:08 +02:00
theorem AnalyticAt.order_eq_zero_iff
{f : }
{z₀ : }
(hf : AnalyticAt f z₀) :
hf.order = 0 ↔ f z₀ ≠ 0 := by
have : (0 : ENat) = (0 : Nat) := by rfl
rw [this, AnalyticAt.order_eq_nat_iff hf 0]
constructor
· intro hz
obtain ⟨g, _, h₂g, h₃g⟩ := hz
simp at h₃g
rw [Filter.Eventually.self_of_nhds h₃g]
tauto
· intro hz
use f
constructor
· exact hf
· constructor
· exact hz
· simp
2024-09-10 14:43:28 +02:00
theorem AnalyticAt.supp_order_toNat
{f : }
{z₀ : }
(hf : AnalyticAt f z₀) :
hf.order.toNat ≠ 0 → f z₀ = 0 := by
contrapose
intro h₁f
simp [hf.order_eq_zero_iff.2 h₁f]
2024-09-11 15:48:43 +02:00
theorem ContinuousLinearEquiv.analyticAt
( : ≃L[] ) (z₀ : ) : AnalyticAt (⇑ℓ) z₀ := .toContinuousLinearMap.analyticAt z₀
theorem eventually_nhds_comp_composition
{f₁ f₂ : }
{z₀ : }
(hf : ∀ᶠ (z : ) in nhds ( z₀), f₁ z = f₂ z)
(h : Continuous ) :
∀ᶠ (z : ) in nhds z₀, (f₁ ∘ ) z = (f₂ ∘ ) z := by
obtain ⟨t, h₁t, h₂t, h₃t⟩ := eventually_nhds_iff.1 hf
apply eventually_nhds_iff.2
use ℓ⁻¹' t
constructor
· intro y hy
exact h₁t ( y) hy
· constructor
· apply IsOpen.preimage
exact ContinuousLinearEquiv.continuous
exact h₂t
· exact h₃t
theorem AnalyticAt.order_congr
{f₁ f₂ : }
{z₀ : }
(hf₁ : AnalyticAt f₁ z₀)
(hf₂ : AnalyticAt f₂ z₀)
(hf : ∀ᶠ (z : ) in nhds z₀, f₁ z = f₂ z) :
hf₁.order = hf₂.order := by
sorry
theorem AnalyticAt.order_comp_CLE
( : ≃L[] )
{f : }
{z₀ : }
(hf : AnalyticAt f ( z₀)) :
hf.order = (hf.comp (.analyticAt z₀)).order := by
by_cases h₁f : hf.order =
· rw [h₁f]
rw [AnalyticAt.order_eq_top_iff] at h₁f
let A := eventually_nhds_comp_composition h₁f .continuous
simp at A
have : AnalyticAt (0 : ) z₀ := by apply analyticAt_const
rw [AnalyticAt.order_congr (hf.comp (.analyticAt z₀)) this A]
have : this.order = := by
rw [AnalyticAt.order_eq_top_iff]
simp
rw [this]
· let n := hf.order.toNat
have hn : hf.order = n := Eq.symm (ENat.coe_toNat h₁f)
rw [hn]
rw [AnalyticAt.order_eq_nat_iff] at hn
obtain ⟨g, h₁g, h₂g, h₃g⟩ := hn
have A := eventually_nhds_comp_composition h₃g .continuous
simp only [Function.comp_apply] at A
have : AnalyticAt (fun z ↦ ( z - z₀) ^ n • g ( z) : ) z₀ := by apply analyticAt_const
rw [AnalyticAt.order_congr (hf.comp (.analyticAt z₀)) this A]
simp
rw [AnalyticAt.order_mul]
--rw [hn, AnalyticAt.order_eq_nat_iff]
rw [AnalyticAt.order_eq_nat_iff] at hn
obtain ⟨g, h₁g, h₂g, h₃g⟩ := hn
use g ∘
constructor
· exact h₁g.comp (.analyticAt z₀)
· constructor
· exact h₂g
· rw [eventually_nhds_iff]
rw [eventually_nhds_iff] at h₃g
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₃g
use ℓ⁻¹' t
constructor
· intro y hy
simp
rw [h₁t ( y) hy]
sorry
· constructor
· apply IsOpen.preimage
exact ContinuousLinearEquiv.continuous
exact h₂t
· exact h₃t