nevanlinna/Nevanlinna/stronglyMeromorphic.lean

212 lines
6.3 KiB
Plaintext
Raw Normal View History

2024-10-08 09:35:17 +02:00
import Mathlib.Analysis.Analytic.Meromorphic
import Nevanlinna.analyticAt
2024-10-09 12:13:22 +02:00
import Nevanlinna.mathlibAddOn
2024-10-08 09:35:17 +02:00
2024-10-08 15:39:19 +02:00
2024-10-23 13:03:41 +02:00
open Topology
2024-10-08 15:39:19 +02:00
/- Strongly MeromorphicAt -/
2024-10-08 09:35:17 +02:00
def StronglyMeromorphicAt
(f : )
(z₀ : ) :=
2024-10-08 15:39:19 +02:00
(∀ᶠ (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))
2024-10-08 09:35:17 +02:00
2024-10-08 15:39:19 +02:00
/- 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
2024-10-09 12:13:22 +02:00
· 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
2024-10-08 15:39:19 +02:00
2024-10-09 06:33:14 +02:00
2024-10-09 12:13:22 +02:00
/- Strongly MeromorphicAt of non-negative order is analytic -/
2024-10-08 15:39:19 +02:00
theorem StronglyMeromorphicAt.analytic
{f : }
{z₀ : }
(h₁f : StronglyMeromorphicAt f z₀)
(h₂f : 0 ≤ h₁f.meromorphicAt.order):
AnalyticAt f z₀ := by
2024-10-09 12:13:22 +02:00
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
2024-10-08 15:39:19 +02:00
2024-10-14 13:41:05 +02:00
/- 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
2024-10-08 15:39:19 +02:00
/- Make strongly MeromorphicAt -/
2024-10-14 13:25:49 +02:00
noncomputable def MeromorphicAt.makeStronglyMeromorphicAt
2024-10-08 09:35:17 +02:00
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
:= by
2024-10-21 10:36:48 +02:00
intro z
by_cases z = z₀
2024-10-22 17:12:59 +02:00
· by_cases h₁f : hf.order = (0 : )
· rw [hf.order_eq_int_iff] at h₁f
exact (Classical.choose h₁f) z₀
2024-10-21 10:36:48 +02:00
· exact 0
· exact f z
2024-10-08 09:35:17 +02:00
2024-10-22 17:12:59 +02:00
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₀) :
2024-10-23 13:03:41 +02:00
f =ᶠ[𝓝[≠] z₀] hf.makeStronglyMeromorphicAt := by
2024-10-22 17:12:59 +02:00
apply eventually_nhdsWithin_of_forall
exact fun x a => m₁ hf x a
lemma Mnhds
{f g : }
{z₀ : }
2024-10-23 13:03:41 +02:00
(h₁ : f =ᶠ[𝓝[≠] z₀] g)
2024-10-22 17:12:59 +02:00
(h₂ : f z₀ = g z₀) :
2024-10-23 13:03:41 +02:00
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
2024-10-08 15:39:19 +02:00
2024-10-24 13:12:03 +02:00
theorem localIdentity
{f g : }
{z₀ : }
(hf : AnalyticAt f z₀)
(hg : AnalyticAt g z₀) :
f =ᶠ[𝓝[≠] z₀] g → f =ᶠ[𝓝 z₀] g := by
2024-10-24 13:46:27 +02:00
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
2024-10-24 13:12:03 +02:00
2024-10-08 15:39:19 +02:00
theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic
2024-10-08 09:35:17 +02:00
{f : }
{z₀ : }
(hf : MeromorphicAt f z₀) :
2024-10-21 10:36:48 +02:00
StronglyMeromorphicAt hf.makeStronglyMeromorphicAt z₀ := by
2024-10-22 17:12:59 +02:00
by_cases h₂f : hf.order =
2024-10-23 13:03:41 +02:00
· have : hf.makeStronglyMeromorphicAt =ᶠ[𝓝 z₀] 0 := by
2024-10-22 17:12:59 +02:00
apply Mnhds
2024-10-23 13:03:41 +02:00
· 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]
2024-10-22 17:12:59 +02:00
apply AnalyticAt.stronglyMeromorphicAt
rw [analyticAt_congr this]
apply analyticAt_const
2024-10-23 13:03:41 +02:00
· 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
2024-10-23 13:31:22 +02:00
by_cases h₃f : hf.order = (0 : )
· let h₄f := (hf.order_eq_int_iff 0).1 h₃f
simp [h₃f]
2024-10-24 12:50:08 +02:00
obtain ⟨h₁G, h₂G, h₃G⟩ := Classical.choose_spec h₄f
simp at h₃G
2024-10-24 13:12:03 +02:00
have hn : n = 0 := Eq.symm ((fun {α} {a} {b} h => (WithTop.eq_untop_iff h).mpr) h₂f (id (Eq.symm h₃f)))
2024-10-24 12:50:08 +02:00
rw [hn]
2024-10-24 13:12:03 +02:00
rw [hn] at h₃g; simp at h₃g
2024-10-24 12:50:08 +02:00
simp
2024-10-24 13:46:27 +02:00
have : g =ᶠ[𝓝 z₀] (Classical.choose h₄f) := by
2024-10-24 13:12:03 +02:00
apply localIdentity h₁g h₁G
exact Filter.EventuallyEq.trans (Filter.EventuallyEq.symm h₃g) h₃G
2024-10-24 12:50:08 +02:00
rw [Filter.EventuallyEq.eq_of_nhds this]
· have : hf.order ≠ 0 := h₃f
simp [this]
2024-10-23 13:03:41 +02:00
left
apply zero_zpow n
dsimp [n]
rwa [WithTop.untop_eq_iff h₂f]