Update stronglyMeromorphic.lean

This commit is contained in:
Stefan Kebekus 2024-10-21 10:36:48 +02:00
parent 8408172272
commit 25b0ffd899
1 changed files with 49 additions and 6 deletions

View File

@ -81,26 +81,69 @@ theorem AnalyticAt.stronglyMeromorphicAt
tauto
theorem MeromorphicAt.order_neq_top_iff
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
hf.order ≠ ↔ ∃ (g : ), AnalyticAt g z₀ ∧ g z₀ ≠ 0 ∧ ∀ᶠ (z : ) in nhdsWithin z₀ {z₀}ᶜ, f z = (z - z₀) ^ (hf.order.untop' 0) • g z := by
rw [← hf.order_eq_int_iff (hf.order.untop' 0)]
constructor
· intro h₁f
apply?
exact Eq.symm (ENat.coe_toNat h₁f)
· intro h₁f
exact ENat.coe_toNat_eq_self.mp (id (Eq.symm h₁f))
/- Make strongly MeromorphicAt -/
noncomputable def MeromorphicAt.makeStronglyMeromorphicAt
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
:= by
by_cases h₂f : hf.order = 0
by_cases h₁f : hf.order =
· exact 0
·
intro z
by_cases z = z₀
· by_cases h₂f : hf.order = 0
· have : (0 : WithTop ) = (0 : ) := rfl
rw [this, hf.order_eq_int_iff] at h₂f
exact Classical.choose h₂f
exact (Classical.choose h₂f) z₀
· exact 0
· exact f z
theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
StronglyMeromorphicAt hf.makeStronglyMeromorphic z₀ := by
StronglyMeromorphicAt hf.makeStronglyMeromorphicAt z₀ := by
by_cases h₂f : hf.order = 0
· apply AnalyticAt.stronglyMeromorphicAt
let A := h₂f
rw [(by rfl : (0 : WithTop ) = (0 : )), hf.order_eq_int_iff] at A
simp at A
have : hf.makeStronglyMeromorphicAt = Classical.choose A := by
simp [MeromorphicAt.makeStronglyMeromorphicAt, h₂f]
let B := Classical.choose_spec A
rw [this]
tauto
· by_cases h₃f : hf.order =
· rw [MeromorphicAt.order_eq_top_iff] at h₃f
left
sorry
· sorry
theorem makeStronglyMeromorphic_eventuallyEq