Working…

This commit is contained in:
Stefan Kebekus 2024-12-21 07:04:13 +01:00
parent 402128875f
commit e488068a7c
2 changed files with 92 additions and 0 deletions

View File

@ -38,6 +38,7 @@ noncomputable def MeromorphicOn.N_infty
:= :=
fun r ↦ ∑ᶠ z, (max 0 (-((hf.restrict |r|).divisor z))) * log (r * ‖z‖⁻¹) fun r ↦ ∑ᶠ z, (max 0 (-((hf.restrict |r|).divisor z))) * log (r * ‖z‖⁻¹)
theorem Nevanlinna_counting₀ theorem Nevanlinna_counting₀
{f : } {f : }
(hf : MeromorphicOn f ) : (hf : MeromorphicOn f ) :

View File

@ -216,5 +216,96 @@ theorem MeromorphicAt.order_inv
· exact ⟨h₂t, h₃t⟩ · exact ⟨h₂t, h₃t⟩
theorem AnalyticAt.meromorphicAt_order_nonneg
{f : }
{z₀ : }
(hf : AnalyticAt f z₀) :
0 ≤ hf.meromorphicAt.order := by
rw [hf.meromorphicAt_order]
rw [(by rfl : (0 : WithTop ) = WithTop.map Nat.cast (0 : ℕ∞))]
rw [WithTop.map_le_iff]
simp; simp
theorem MeromorphicAt.order_add
{f₁ f₂ : }
{z₀ : }
(hf₁ : MeromorphicAt f₁ z₀)
(hf₂ : MeromorphicAt f₂ z₀) :
(hf₁.add hf₂).order ≤ min hf₁.order hf₂.order := by
-- Handle the trivial cases where one of the orders equals
by_cases h₂f₁: hf₁.order =
· rw [h₂f₁]; simp
rw [hf₁.order_eq_top_iff] at h₂f₁
have h : f₁ + f₂ =ᶠ[𝓝[≠] z₀] f₂ := by
rw [eventuallyEq_nhdsWithin_iff, eventually_iff_exists_mem]
rw [eventually_nhdsWithin_iff, eventually_iff_exists_mem] at h₂f₁
obtain ⟨v, hv⟩ := h₂f₁
use v; simp; trivial
rw [(hf₁.add hf₂).order_congr h]
by_cases h₂f₂: hf₂.order =
· rw [h₂f₂]; simp
rw [hf₂.order_eq_top_iff] at h₂f₂
have h : f₁ + f₂ =ᶠ[𝓝[≠] z₀] f₁ := by
rw [eventuallyEq_nhdsWithin_iff, eventually_iff_exists_mem]
rw [eventually_nhdsWithin_iff, eventually_iff_exists_mem] at h₂f₂
obtain ⟨v, hv⟩ := h₂f₂
use v; simp; trivial
rw [(hf₁.add hf₂).order_congr h]
obtain ⟨g₁, h₁g₁, h₂g₁, h₃g₁⟩ := hf₁.order_neg_zero_iff.1 h₂f₁
obtain ⟨g₂, h₁g₂, h₂g₂, h₃g₂⟩ := hf₂.order_neg_zero_iff.1 h₂f₂
let n₁ := WithTop.untop' 0 hf₁.order
let n₂ := WithTop.untop' 0 hf₁.order
let n := min n₁ n₂
have h₁n₁ : 0 ≤ n₁ - n := by
rw [sub_nonneg]
exact Int.min_le_left n₁ n₂
have h₁n₂ : 0 ≤ n₂ - n := by
rw [sub_nonneg]
exact Int.min_le_right n₁ n₂
let g := (fun z ↦ (z - z₀) ^ (n₁ - n)) * g₁ + (fun z ↦ (z - z₀) ^ (n₂ - n)) * g₂
have h₁g : AnalyticAt g z₀ := by
apply AnalyticAt.add
apply AnalyticAt.mul
apply AnalyticAt.zpow_nonneg _ h₁n₁
apply AnalyticAt.sub
apply analyticAt_id
apply analyticAt_const
exact h₁g₁
apply AnalyticAt.mul
apply AnalyticAt.zpow_nonneg _ h₁n₁
apply AnalyticAt.sub
apply analyticAt_id
apply analyticAt_const
exact h₁g₂
have h₂g : 0 ≤ h₁g.meromorphicAt.order := h₁g.meromorphicAt_order_nonneg
have : f₁ + f₂ =ᶠ[𝓝[≠] z₀] (fun z ↦ (z - z₀) ^ n) * g := by
sorry
rw [(hf₁.add hf₂).order_congr this]
have t₀ : MeromorphicAt (fun z ↦ (z - z₀) ^ n) z₀ := by
sorry
rw [t₀.order_mul h₁g.meromorphicAt]
have t₁ : t₀.order = n := by
sorry
rw [t₁]
-- Exercise in WithTop
sorry
theorem MeromorphicAt.order_add_of_ne_orders
{f₁ f₂ : }
{z₀ : }
(hf₁ : MeromorphicAt f₁ z₀)
(hf₂ : MeromorphicAt f₂ z₀)
(hf₁₂ : hf₁.order ≠ hf₂.order) :
(hf₁.add hf₂).order ≤ hf₁.order + hf₂.order := by
sorry
-- might want theorem MeromorphicAt.order_zpow -- might want theorem MeromorphicAt.order_zpow