nevanlinna/Nevanlinna/analyticAt.lean

200 lines
5.9 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Analysis.Analytic.IsolatedZeros
import Mathlib.Analysis.Complex.Basic
import Mathlib.Analysis.Analytic.Linear
theorem AnalyticAt.order_mul
{f₁ f₂ : }
{z₀ : }
(hf₁ : AnalyticAt f₁ z₀)
(hf₂ : AnalyticAt f₂ z₀) :
(hf₁.mul 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₂
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
theorem AnalyticAt.order_pow
{f : }
{z₀ : }
{n : }
(hf : AnalyticAt f z₀) :
(hf.pow n).order = n * hf.order := by
induction' n with n hn
· simp; rw [AnalyticAt.order_eq_zero_iff]; simp
· simp
simp_rw [add_mul, pow_add]
simp
rw [AnalyticAt.order_mul (hf.pow n) hf]
rw [hn]
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]
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
have : AnalyticAt (fun z ↦ ( z - z₀) ^ n • g ( z) : ) z₀ := by
sorry
rw [AnalyticAt.order_congr (hf.comp (.analyticAt z₀)) this A]
simp
have t₀ : AnalyticAt (fun z => ( z - z₀) ^ n) z₀ := by
sorry
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
rw [AnalyticAt.order_eq_nat_iff]
use (fun z ↦ 1)
simp
constructor
· exact analyticAt_const
· apply Filter.Eventually.of_forall
intro x
sorry
have : t₀.order = n := by
rw [AnalyticAt.order_pow t₁, this]
simp
rw [this]
have : (comp h₁g (ContinuousLinearEquiv.analyticAt z₀)).order = 0 := by
rwa [AnalyticAt.order_eq_zero_iff]
rw [this]
simp