nevanlinna/Nevanlinna/holomorphic_zero.lean

390 lines
10 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 Init.Classical
import Mathlib.Analysis.Analytic.Meromorphic
import Mathlib.Topology.ContinuousOn
import Mathlib.Analysis.Analytic.IsolatedZeros
import Nevanlinna.holomorphic
noncomputable def zeroDivisor
(f : ) :
:= by
intro z
by_cases hf : AnalyticAt f z
· exact hf.order.toNat
· exact 0
theorem analyticAtZeroDivisorSupport
{f : }
{z : }
(h : z ∈ Function.support (zeroDivisor f)) :
AnalyticAt f z := by
by_contra h₁f
simp at h
dsimp [zeroDivisor] at h
simp [h₁f] at h
theorem zeroDivisor_eq_ord_AtZeroDivisorSupport
{f : }
{z : }
(h : z ∈ Function.support (zeroDivisor f)) :
zeroDivisor f z = (analyticAtZeroDivisorSupport h).order.toNat := by
unfold zeroDivisor
simp [analyticAtZeroDivisorSupport h]
lemma toNatEqSelf_iff {n : ℕ∞} : n.toNat = n ↔ ∃ m : , m = n := by
constructor
· intro H₁
rw [← ENat.some_eq_coe, ← WithTop.ne_top_iff_exists]
by_contra H₂
rw [H₂] at H₁
simp at H₁
· intro H
obtain ⟨m, hm⟩ := H
rw [← hm]
simp
lemma natural_if_toNatNeZero {n : ℕ∞} : n.toNat ≠ 0 → ∃ m : , m = n := by
rw [← ENat.some_eq_coe, ← WithTop.ne_top_iff_exists]
contrapose; simp; tauto
theorem zeroDivisor_localDescription
{f : }
{z₀ : }
(h : z₀ ∈ Function.support (zeroDivisor f)) :
∃ (g : ), AnalyticAt g z₀ ∧ g z₀ ≠ 0 ∧ ∀ᶠ (z : ) in nhds z₀, f z = (z - z₀) ^ (zeroDivisor f z₀) • g z := by
have A : zeroDivisor f ↑z₀ ≠ 0 := by exact h
let B := zeroDivisor_eq_ord_AtZeroDivisorSupport h
rw [B] at A
have C := natural_if_toNatNeZero A
obtain ⟨m, hm⟩ := C
have h₂m : m ≠ 0 := by
rw [← hm] at A
simp at A
assumption
rw [eq_comm] at hm
let E := AnalyticAt.order_eq_nat_iff (analyticAtZeroDivisorSupport h) m
let F := hm
rw [E] at F
have : m = zeroDivisor f z₀ := by
rw [B, hm]
simp
rwa [this] at F
theorem zeroDivisor_zeroSet
{f : }
{z₀ : }
(h : z₀ ∈ Function.support (zeroDivisor f)) :
f z₀ = 0 := by
obtain ⟨g, _, _, h₃⟩ := zeroDivisor_localDescription h
rw [Filter.Eventually.self_of_nhds h₃]
simp
left
exact h
theorem zeroDivisor_support_iff
{f : }
{z₀ : } :
z₀ ∈ Function.support (zeroDivisor f) ↔
f z₀ = 0 ∧
AnalyticAt f z₀ ∧
∃ (g : ), AnalyticAt g z₀ ∧ g z₀ ≠ 0 ∧ ∀ᶠ (z : ) in nhds z₀, f z = (z - z₀) ^ (zeroDivisor f z₀) • g z := by
constructor
· intro hz
constructor
· exact zeroDivisor_zeroSet hz
· constructor
· exact analyticAtZeroDivisorSupport hz
· exact zeroDivisor_localDescription hz
· intro ⟨h₁, h₂, h₃⟩
have : zeroDivisor f z₀ = (h₂.order).toNat := by
unfold zeroDivisor
simp [h₂]
simp [this]
simp [(h₂.order_eq_nat_iff (zeroDivisor f z₀)).2 h₃]
obtain ⟨g, h₁g, h₂g, h₃g⟩ := h₃
rw [Filter.Eventually.self_of_nhds h₃g] at h₁
simp [h₂g] at h₁
assumption
theorem topOnPreconnected
{f : }
{U : Set }
(hU : IsPreconnected U)
(h₁f : AnalyticOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) :
∀ (hz : z ∈ U), (h₁f z hz).order ≠ := by
by_contra H
push_neg at H
obtain ⟨z', hz'⟩ := H
rw [AnalyticAt.order_eq_top_iff] at hz'
rw [← AnalyticAt.frequently_zero_iff_eventually_zero (h₁f z z')] at hz'
have A := AnalyticOn.eqOn_zero_of_preconnected_of_frequently_eq_zero h₁f hU z' hz'
tauto
theorem supportZeroSet
{f : }
{U : Set }
(hU : IsPreconnected U)
(h₁f : AnalyticOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) :
U ∩ Function.support (zeroDivisor f) = U ∩ f⁻¹' {0} := by
ext x
constructor
· intro hx
constructor
· exact hx.1
exact zeroDivisor_zeroSet hx.2
· simp
intro h₁x h₂x
constructor
· exact h₁x
· let A := zeroDivisor_support_iff (f := f) (z₀ := x)
simp at A
rw [A]
constructor
· exact h₂x
· constructor
· exact h₁f x h₁x
· have B := AnalyticAt.order_eq_nat_iff (h₁f x h₁x) (zeroDivisor f x)
simp at B
rw [← B]
dsimp [zeroDivisor]
simp [h₁f x h₁x]
refine Eq.symm (ENat.coe_toNat ?h.mpr.right.right.right.a)
exact topOnPreconnected hU h₁f h₂f h₁x
theorem discreteZeros
{f : } :
DiscreteTopology (Function.support (zeroDivisor f)) := by
simp_rw [← singletons_open_iff_discrete, Metric.isOpen_singleton_iff]
intro z
have A : zeroDivisor f ↑z ≠ 0 := by exact z.2
let B := zeroDivisor_eq_ord_AtZeroDivisorSupport z.2
rw [B] at A
have C := natural_if_toNatNeZero A
obtain ⟨m, hm⟩ := C
have h₂m : m ≠ 0 := by
rw [← hm] at A
simp at A
assumption
rw [eq_comm] at hm
let E := AnalyticAt.order_eq_nat_iff (analyticAtZeroDivisorSupport z.2) m
rw [E] at hm
obtain ⟨g, h₁g, h₂g, h₃g⟩ := hm
rw [Metric.eventually_nhds_iff_ball] at h₃g
have : ∃ ε > 0, ∀ y ∈ Metric.ball (↑z) ε, g y ≠ 0 := by
have h₄g : ContinuousAt g z := AnalyticAt.continuousAt h₁g
have : {0}ᶜ ∈ nhds (g z) := by
exact compl_singleton_mem_nhds_iff.mpr h₂g
let F := h₄g.preimage_mem_nhds this
rw [Metric.mem_nhds_iff] at F
obtain ⟨ε, h₁ε, h₂ε⟩ := F
use ε
constructor; exact h₁ε
intro y hy
let G := h₂ε hy
simp at G
exact G
obtain ⟨ε₁, h₁ε₁⟩ := this
obtain ⟨ε₂, h₁ε₂, h₂ε₂⟩ := h₃g
use min ε₁ ε₂
constructor
· have : 0 < min ε₁ ε₂ := by
rw [lt_min_iff]
exact And.imp_right (fun _ => h₁ε₂) h₁ε₁
exact this
intro y
intro h₁y
have h₂y : ↑y ∈ Metric.ball (↑z) ε₂ := by
simp
calc dist y z
_ < min ε₁ ε₂ := by assumption
_ ≤ ε₂ := by exact min_le_right ε₁ ε₂
have h₃y : ↑y ∈ Metric.ball (↑z) ε₁ := by
simp
calc dist y z
_ < min ε₁ ε₂ := by assumption
_ ≤ ε₁ := by exact min_le_left ε₁ ε₂
let F := h₂ε₂ y.1 h₂y
rw [zeroDivisor_zeroSet y.2] at F
simp at F
simp [h₂m] at F
have : g y.1 ≠ 0 := by
exact h₁ε₁.2 y h₃y
simp [this] at F
ext
rwa [sub_eq_zero] at F
theorem zeroDivisor_finiteOnCompact
{f : }
{U : Set }
(hU : IsPreconnected U)
(h₁f : AnalyticOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) -- not needed!
(h₂U : IsCompact U) :
Set.Finite (U ∩ Function.support (zeroDivisor f)) := by
have hinter : IsCompact (U ∩ Function.support (zeroDivisor f)) := by
apply IsCompact.of_isClosed_subset h₂U
rw [supportZeroSet]
apply h₁f.continuousOn.preimage_isClosed_of_isClosed
exact IsCompact.isClosed h₂U
exact isClosed_singleton
assumption
assumption
assumption
exact Set.inter_subset_left
apply hinter.finite
apply DiscreteTopology.of_subset (s := Function.support (zeroDivisor f))
exact discreteZeros (f := f)
exact Set.inter_subset_right
theorem AnalyticOn.order_eq_nat_iff
{f : }
{U : Set }
{z₀ : }
(hf : AnalyticOn f U)
(hz₀ : z₀ ∈ U)
(n : ) :
(hf z₀ hz₀).order = ↑n → ∃ (g : ), AnalyticOn g U ∧ g z₀ ≠ 0 ∧ ∀ z ∈ U, f z = (z - z₀) ^ n • g z := by
intro hn
obtain ⟨gloc, h₁gloc, h₂gloc, h₃gloc⟩ := (AnalyticAt.order_eq_nat_iff (hf z₀ hz₀) n).1 hn
let g : := fun z ↦ if hz : z = z₀ then gloc z₀ else (f z) / (z - z₀) ^ n
have t₀ : ∀ᶠ (z : ) in nhds z₀, g z = gloc z := by
rw [eventually_nhds_iff]
rw [eventually_nhds_iff] at h₃gloc
obtain ⟨t, h₁t, h₂t, h₃t⟩ := h₃gloc
use t
constructor
· intro y h₁y
by_cases h₂y : y = z₀
· dsimp [g]; simp [h₂y]
· dsimp [g]; simp [h₂y]
rw [div_eq_iff_mul_eq, eq_comm, mul_comm]
exact h₁t y h₁y
norm_num
rw [sub_eq_zero]
tauto
· constructor
· assumption
· assumption
have t₁ {z₁ : } : z₁ ≠ z₀ → ∀ᶠ (z : ) in nhds z₁, g z = f z / (z - z₀) ^ n := by
intro hz₁
rw [eventually_nhds_iff]
use {z₀}ᶜ
constructor
· intro y hy
simp at hy
dsimp [g]
simp [hy]
· constructor
· exact isOpen_compl_singleton
· tauto
use g
constructor
· -- AnalyticOn g U
intro z h₁z
by_cases h₂z : z = z₀
· rw [h₂z]
apply AnalyticAt.congr h₁gloc
exact Filter.EventuallyEq.symm t₀
· simp_rw [eq_comm] at t₁
apply AnalyticAt.congr _ (t₁ h₂z)
apply AnalyticAt.div
exact hf z h₁z
apply AnalyticAt.pow
apply AnalyticAt.sub
apply analyticAt_id
exact analyticAt_const
simp
rw [sub_eq_zero]
tauto
theorem eliminatingZeros
{f : }
{U : Set }
(h₁U : IsPreconnected U)
(h₂U : IsCompact U)
(h₁f : AnalyticOn f U)
(h₂f : ∃ z ∈ U, f z ≠ 0) :
∃ F : , (AnalyticOn F U) ∧ (f = F * ∏ᶠ a ∈ (U ∩ (zeroDivisor f).support), fun z ↦ (z - a) ^ (zeroDivisor f a)) := by
have hs := zeroDivisor_finiteOnCompact h₁U h₁f h₂f h₂U
rw [finprod_mem_eq_finite_toFinset_prod _ hs]
let A := hs.card_eq
let F : := by
intro z
if hz : z ∈ U ∩ (zeroDivisor f).support then
exact (Classical.choose (zeroDivisor_support_iff.1 hz.2).2.2) z
else
exact (f z) / ∏ i ∈ hs.toFinset, (z - i) ^ zeroDivisor f i
use F
constructor
· intro z h₁z
by_cases h₂z : z ∈ (zeroDivisor f).support
· -- case: z ∈ Function.support (zeroDivisor f)
have : z ∈ U ∩ (zeroDivisor f).support := by exact Set.mem_inter h₁z h₂z
sorry
· sorry
have : MeromorphicOn F U := by
apply MeromorphicOn.div
exact AnalyticOn.meromorphicOn h₁f
apply AnalyticOn.meromorphicOn
apply Finset.analyticOn_prod
intro n hn
apply AnalyticOn.pow
apply AnalyticOn.sub
exact analyticOn_id
exact analyticOn_const
--use F
sorry