Update stronglyMeromorphic.lean

This commit is contained in:
Stefan Kebekus 2024-10-24 13:46:27 +02:00
parent 6aa5bc3b1c
commit f373bf786b
1 changed files with 15 additions and 3 deletions

View File

@ -140,7 +140,20 @@ theorem localIdentity
(hf : AnalyticAt f z₀) (hf : AnalyticAt f z₀)
(hg : AnalyticAt g z₀) : (hg : AnalyticAt g z₀) :
f =ᶠ[𝓝[≠] z₀] g → f =ᶠ[𝓝 z₀] g := by f =ᶠ[𝓝[≠] z₀] g → f =ᶠ[𝓝 z₀] g := by
sorry intro h
let Δ := f - g
have : AnalyticAt Δ z₀ := AnalyticAt.sub hf hg
have t₁ : Δ =ᶠ[𝓝[≠] z₀] 0 := by
exact Filter.eventuallyEq_iff_sub.mp h
have : Δ =ᶠ[𝓝 z₀] 0 := by
rcases (AnalyticAt.eventually_eq_zero_or_eventually_ne_zero this) with h | h
· exact h
· have := Filter.EventuallyEq.eventually t₁
let A := Filter.eventually_and.2 ⟨this, h⟩
let _ := Filter.Eventually.exists A
tauto
exact Filter.eventuallyEq_iff_sub.mpr this
theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
@ -179,7 +192,6 @@ theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
simp simp
by_cases h₃f : hf.order = (0 : ) by_cases h₃f : hf.order = (0 : )
· let h₄f := (hf.order_eq_int_iff 0).1 h₃f · let h₄f := (hf.order_eq_int_iff 0).1 h₃f
let G := Classical.choose h₄f
simp [h₃f] simp [h₃f]
obtain ⟨h₁G, h₂G, h₃G⟩ := Classical.choose_spec h₄f obtain ⟨h₁G, h₂G, h₃G⟩ := Classical.choose_spec h₄f
simp at h₃G simp at h₃G
@ -187,7 +199,7 @@ theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
rw [hn] rw [hn]
rw [hn] at h₃g; simp at h₃g rw [hn] at h₃g; simp at h₃g
simp simp
have : g =ᶠ[𝓝 z₀] G := by have : g =ᶠ[𝓝 z₀] (Classical.choose h₄f) := by
apply localIdentity h₁g h₁G apply localIdentity h₁g h₁G
exact Filter.EventuallyEq.trans (Filter.EventuallyEq.symm h₃g) h₃G exact Filter.EventuallyEq.trans (Filter.EventuallyEq.symm h₃g) h₃G
rw [Filter.EventuallyEq.eq_of_nhds this] rw [Filter.EventuallyEq.eq_of_nhds this]