2024-09-14 08:38:04 +02:00
|
|
|
|
import Mathlib.Analysis.Analytic.Meromorphic
|
2024-09-13 09:21:57 +02:00
|
|
|
|
import Nevanlinna.holomorphic_JensenFormula
|
|
|
|
|
|
|
|
|
|
open Real
|
|
|
|
|
|
2024-09-14 08:38:04 +02:00
|
|
|
|
variable (f : ℂ → ℂ)
|
|
|
|
|
variable {h₁f : MeromorphicOn f ⊤}
|
2024-09-13 09:21:57 +02:00
|
|
|
|
variable {h₂f : f 0 ≠ 0}
|
|
|
|
|
|
2024-09-14 08:38:04 +02:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
noncomputable def Nevanlinna.n : ℝ → ℤ := by
|
2024-09-13 09:21:57 +02:00
|
|
|
|
intro r
|
2024-09-14 08:38:04 +02:00
|
|
|
|
have : MeromorphicAt f z := by
|
|
|
|
|
exact h₁f (sorryAx ℂ true) trivial
|
|
|
|
|
exact ∑ᶠ z ∈ Metric.ball (0 : ℂ) r, (h₁f z trivial).order
|
2024-09-13 09:21:57 +02:00
|
|
|
|
|
|
|
|
|
noncomputable def integratedCounting : ℝ → ℝ := by
|
|
|
|
|
intro r
|
|
|
|
|
-- Issue here: for s on the boundary of the ball, zero comes out.
|
|
|
|
|
exact ∑ᶠ s, (h₁f.order s).toNat * log (r * ‖s.1‖⁻¹)
|