nevanlinna/Nevanlinna/stronglyMeromorphic.lean

180 lines
5.0 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.Meromorphic
import Nevanlinna.analyticAt
import Nevanlinna.mathlibAddOn
open Topology
/- Strongly MeromorphicAt -/
def StronglyMeromorphicAt
(f : )
(z₀ : ) :=
(∀ᶠ (z : ) in nhds z₀, f z = 0) (∃ (n : ), ∃ g : , (AnalyticAt g z₀) ∧ (g z₀ ≠ 0) ∧ (∀ᶠ (z : ) in nhds z₀, f z = (z - z₀) ^ n • g z))
/- Strongly MeromorphicAt is Meromorphic -/
theorem StronglyMeromorphicAt.meromorphicAt
{f : }
{z₀ : }
(hf : StronglyMeromorphicAt f z₀) :
MeromorphicAt f z₀ := by
rcases hf with h|h
· use 0; simp
rw [analyticAt_congr h]
exact analyticAt_const
· obtain ⟨n, g, h₁g, _, h₃g⟩ := h
rw [meromorphicAt_congr' h₃g]
apply MeromorphicAt.smul
apply MeromorphicAt.zpow
apply MeromorphicAt.sub
apply MeromorphicAt.id
apply MeromorphicAt.const
exact AnalyticAt.meromorphicAt h₁g
/- Strongly MeromorphicAt of non-negative order is analytic -/
theorem StronglyMeromorphicAt.analytic
{f : }
{z₀ : }
(h₁f : StronglyMeromorphicAt f z₀)
(h₂f : 0 ≤ h₁f.meromorphicAt.order):
AnalyticAt f z₀ := by
let h₁f' := h₁f
rcases h₁f' with h|h
· rw [analyticAt_congr h]
exact analyticAt_const
· obtain ⟨n, g, h₁g, h₂g, h₃g⟩ := h
rw [analyticAt_congr h₃g]
have : h₁f.meromorphicAt.order = n := by
rw [MeromorphicAt.order_eq_int_iff]
use g
constructor
· exact h₁g
· constructor
· exact h₂g
· exact Filter.EventuallyEq.filter_mono h₃g nhdsWithin_le_nhds
rw [this] at h₂f
apply AnalyticAt.smul
nth_rw 1 [← Int.toNat_of_nonneg (WithTop.coe_nonneg.mp h₂f)]
apply AnalyticAt.pow
apply AnalyticAt.sub
apply analyticAt_id -- Warning: want apply AnalyticAt.id
apply analyticAt_const -- Warning: want AnalyticAt.const
exact h₁g
/- Analytic functions are strongly meromorphic -/
theorem AnalyticAt.stronglyMeromorphicAt
{f : }
{z₀ : }
(h₁f : AnalyticAt f z₀) :
StronglyMeromorphicAt f z₀ := by
by_cases h₂f : h₁f.order =
· rw [AnalyticAt.order_eq_top_iff] at h₂f
tauto
· have : h₁f.order ≠ := h₂f
rw [← ENat.coe_toNat_eq_self] at this
rw [eq_comm, AnalyticAt.order_eq_nat_iff] at this
right
use h₁f.order.toNat
obtain ⟨g, h₁g, h₂g, h₃g⟩ := this
use g
tauto
/- Make strongly MeromorphicAt -/
noncomputable def MeromorphicAt.makeStronglyMeromorphicAt
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
:= by
intro z
by_cases z = z₀
· by_cases h₁f : hf.order = (0 : )
· rw [hf.order_eq_int_iff] at h₁f
exact (Classical.choose h₁f) z₀
· exact 0
· exact f z
lemma m₁
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
∀ z ≠ z₀, f z = hf.makeStronglyMeromorphicAt z := by
intro z hz
unfold MeromorphicAt.makeStronglyMeromorphicAt
simp [hz]
lemma m₂
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
f =ᶠ[𝓝[≠] z₀] hf.makeStronglyMeromorphicAt := by
apply eventually_nhdsWithin_of_forall
exact fun x a => m₁ hf x a
lemma Mnhds
{f g : }
{z₀ : }
(h₁ : f =ᶠ[𝓝[≠] z₀] g)
(h₂ : f z₀ = g z₀) :
f =ᶠ[𝓝 z₀] g := by
apply eventually_nhds_iff.2
obtain ⟨t, h₁t, h₂t⟩ := eventually_nhds_iff.1 (eventually_nhdsWithin_iff.1 h₁)
use t
constructor
· intro y hy
by_cases h₂y : y ∈ ({z₀}ᶜ : Set )
· exact h₁t y hy h₂y
· simp at h₂y
rwa [h₂y]
· exact h₂t
theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
StronglyMeromorphicAt hf.makeStronglyMeromorphicAt z₀ := by
by_cases h₂f : hf.order =
· have : hf.makeStronglyMeromorphicAt =ᶠ[𝓝 z₀] 0 := by
apply Mnhds
· apply Filter.EventuallyEq.trans (Filter.EventuallyEq.symm (m₂ hf))
exact (MeromorphicAt.order_eq_top_iff hf).1 h₂f
· unfold MeromorphicAt.makeStronglyMeromorphicAt
simp [h₂f]
apply AnalyticAt.stronglyMeromorphicAt
rw [analyticAt_congr this]
apply analyticAt_const
· let n := hf.order.untop h₂f
have : hf.order = n := by
exact Eq.symm (WithTop.coe_untop hf.order h₂f)
rw [hf.order_eq_int_iff] at this
obtain ⟨g, h₁g, h₂g, h₃g⟩ := this
right
use n
use g
constructor
· assumption
· constructor
· assumption
· apply Mnhds
· apply Filter.EventuallyEq.trans (Filter.EventuallyEq.symm (m₂ hf))
exact h₃g
· unfold MeromorphicAt.makeStronglyMeromorphicAt
simp
by_cases h₃f : hf.order = 0
· simp [h₃f]
sorry
· simp [h₃f]
left
apply zero_zpow n
dsimp [n]
rwa [WithTop.untop_eq_iff h₂f]