Update stronglyMeromorphic.lean

This commit is contained in:
Stefan Kebekus 2024-10-08 15:39:19 +02:00
parent 67b78ad72d
commit 1c31e68e2a
1 changed files with 51 additions and 3 deletions

View File

@ -1,21 +1,69 @@
import Mathlib.Analysis.Analytic.Meromorphic import Mathlib.Analysis.Analytic.Meromorphic
import Nevanlinna.analyticAt import Nevanlinna.analyticAt
/- Strongly MeromorphicAt -/
def StronglyMeromorphicAt def StronglyMeromorphicAt
(f : ) (f : )
(z₀ : ) := (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)) (∀ᶠ (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))
def MeromorphicAt.makeStronglyMeromorphic
/- 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₃g⟩ := h
have : MeromorphicAt (fun z ↦ (z - z₀) ^ n • g z) z₀ := by
simp
apply MeromorphicAt.mul
apply MeromorphicAt.zpow
apply MeromorphicAt.sub
sorry
apply MeromorphicAt.congr this
rw [Filter.eventuallyEq_comm]
exact Filter.EventuallyEq.filter_mono h₃g nhdsWithin_le_nhds
/- Strongly MeromorphicAt of positive order is analytic -/
theorem StronglyMeromorphicAt.analytic
{f : }
{z₀ : }
(h₁f : StronglyMeromorphicAt f z₀)
(h₂f : 0 ≤ h₁f.meromorphicAt.order):
AnalyticAt f z₀ := by
sorry
/- Make strongly MeromorphicAt -/
def MeromorphicAt.makeStronglyMeromorphicAt
{f : } {f : }
{z₀ : } {z₀ : }
(hf : MeromorphicAt f z₀) : (hf : MeromorphicAt f z₀) :
:= by := by
exact 0 exact 0
theorem makeStronglyMeromorphic
theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
{f : } {f : }
{z₀ : } {z₀ : }
(hf : MeromorphicAt f z₀) : (hf : MeromorphicAt f z₀) :
StronglyMeromorphicAt hf.makeStronglyMeromorphic z₀ := by StronglyMeromorphicAt hf.makeStronglyMeromorphic z₀ := by
sorry sorry
theorem makeStronglyMeromorphic_eventuallyEq
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
∀ᶠ (z : ) in nhdsWithin z₀ {z₀}ᶜ, f z = hf.makeStronglyMeromorphicAt z := by
sorry