nevanlinna/Nevanlinna/firstMain.lean
2024-12-24 06:50:23 +01:00

262 lines
6.0 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.MeasureTheory.Integral.CircleIntegral
import Nevanlinna.divisor
import Nevanlinna.meromorphicOn_divisor
import Nevanlinna.meromorphicOn_integrability
import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.stronglyMeromorphic_JensenFormula
open Real
-- Lang p. 164
theorem MeromorphicOn.restrict
{f : }
(h₁f : MeromorphicOn f )
(r : ) :
MeromorphicOn f (Metric.closedBall 0 r) := by
exact fun x a => h₁f x trivial
theorem MeromorphicOn.restrict_inv
{f : }
(h₁f : MeromorphicOn f )
(r : ) :
h₁f.inv.restrict r = (h₁f.restrict r).inv := by
funext x
simp
noncomputable def MeromorphicOn.N_zero
{f : }
(hf : MeromorphicOn f ) :
:=
fun r ↦ ∑ᶠ z, (max 0 ((hf.restrict |r|).divisor z)) * log (r * ‖z‖⁻¹)
noncomputable def MeromorphicOn.N_infty
{f : }
(hf : MeromorphicOn f ) :
:=
fun r ↦ ∑ᶠ z, (max 0 (-((hf.restrict |r|).divisor z))) * log (r * ‖z‖⁻¹)
theorem Nevanlinna_counting₁₁
{f : }
{a : }
(hf : MeromorphicOn f ) :
(hf.add (MeromorphicOn.const a)).N_infty = hf.N_infty := by
have {z : } : 0 < (hf z trivial).order → (hf z trivial).order = ((hf.add (MeromorphicOn.const a)) z trivial).order:= by
intro h
let A := (MeromorphicAt.const a)
rw [←MeromorphicAt.order_add_of_ne_orders (hf z trivial)]
simp
sorry
funext r
unfold MeromorphicOn.N_infty
let A := (hf.restrict |r|).divisor.finiteSupport (isCompact_closedBall 0 |r|)
repeat
rw [finsum_eq_sum_of_support_subset (s := A.toFinset)]
apply Finset.sum_congr rfl
intro x hx
congr 2
simp at hx
sorry
theorem Nevanlinna_counting₀
{f : }
(hf : MeromorphicOn f ) :
hf.inv.N_infty = hf.N_zero := by
funext r
unfold MeromorphicOn.N_zero MeromorphicOn.N_infty
let A := (hf.restrict |r|).divisor.finiteSupport (isCompact_closedBall 0 |r|)
repeat
rw [finsum_eq_sum_of_support_subset (s := A.toFinset)]
apply Finset.sum_congr rfl
intro x hx
congr
let B := hf.restrict_inv |r|
rw [MeromorphicOn.divisor_inv]
simp
--
exact fun x a => hf x trivial
--
intro x
contrapose
simp
intro hx
rw [hx]
tauto
--
intro x
contrapose
simp
intro hx h₁x
rw [MeromorphicOn.divisor_inv (hf.restrict |r|)] at h₁x
simp at h₁x
rw [hx] at h₁x
tauto
theorem Nevanlinna_counting
{f : }
(hf : MeromorphicOn f ) :
hf.N_zero - hf.N_infty = fun r ↦ ∑ᶠ z, ((hf.restrict |r|).divisor z) * log (r * ‖z‖⁻¹) := by
funext r
simp only [Pi.sub_apply]
unfold MeromorphicOn.N_zero MeromorphicOn.N_infty
let A := (hf.restrict |r|).divisor.finiteSupport (isCompact_closedBall 0 |r|)
repeat
rw [finsum_eq_sum_of_support_subset (s := A.toFinset)]
rw [← Finset.sum_sub_distrib]
simp_rw [← sub_mul]
congr
funext x
congr
by_cases h : 0 ≤ (hf.restrict |r|).divisor x
· simp [h]
· have h' : 0 ≤ -((hf.restrict |r|).divisor x) := by
simp at h
apply Int.le_neg_of_le_neg
simp
exact Int.le_of_lt h
simp at h
simp [h']
linarith
--
repeat
intro x
contrapose
simp
intro hx
rw [hx]
tauto
--
noncomputable def MeromorphicOn.m_infty
{f : }
(_ : MeromorphicOn f ) :
:=
fun r ↦ (2 * π)⁻¹ * ∫ x in (0)..(2 * π), logpos ‖f (circleMap 0 r x)‖
theorem Nevanlinna_proximity
{f : }
{r : }
(h₁f : MeromorphicOn f ) :
(2 * π)⁻¹ * ∫ x in (0)..(2 * π), log ‖f (circleMap 0 r x)‖ = (h₁f.m_infty r) - (h₁f.inv.m_infty r) := by
unfold MeromorphicOn.m_infty
rw [← mul_sub]; congr
rw [← intervalIntegral.integral_sub]; congr
funext x
simp_rw [loglogpos]; congr
exact Eq.symm (IsAbsoluteValue.abv_inv Norm.norm (f (circleMap 0 r x)))
--
apply MeromorphicOn.integrable_logpos_abs_f
intro z hx
exact h₁f z trivial
--
apply MeromorphicOn.integrable_logpos_abs_f
exact MeromorphicOn.inv_iff.mpr fun x a => h₁f x trivial
noncomputable def MeromorphicOn.T_infty
{f : }
(hf : MeromorphicOn f ) :
:=
hf.m_infty + hf.N_infty
theorem Nevanlinna_firstMain₁
{f : }
(h₁f : MeromorphicOn f )
(h₂f : StronglyMeromorphicAt f 0)
(h₃f : f 0 ≠ 0) :
(fun _ ↦ log ‖f 0‖) + h₁f.inv.T_infty = h₁f.T_infty := by
rw [add_eq_of_eq_sub]
unfold MeromorphicOn.T_infty
have {A B C D : } : A + B - (C + D) = A - C - (D - B) := by
ring
rw [this]
clear this
rw [Nevanlinna_counting₀ h₁f]
rw [Nevanlinna_counting h₁f]
funext r
simp
rw [← Nevanlinna_proximity h₁f]
by_cases h₁r : r = 0
rw [h₁r]
simp
have : π⁻¹ * 2⁻¹ * (2 * π * log (Complex.abs (f 0))) = (π⁻¹ * (2⁻¹ * 2) * π) * log (Complex.abs (f 0)) := by
ring
rw [this]
clear this
simp [pi_ne_zero]
by_cases hr : 0 < r
let A := jensen hr f (h₁f.restrict r) h₂f h₃f
simp at A
rw [A]
clear A
simp
have {A B : } : -A + B = B - A := by ring
rw [this]
have : |r| = r := by
rw [← abs_of_pos hr]
simp
rw [this]
-- case 0 < -r
have h₂r : 0 < -r := by
simp [h₁r, hr]
by_contra hCon
-- Assume ¬(r < 0), which means r >= 0
push_neg at hCon
-- Now h is r ≥ 0, so we split into cases
rcases lt_or_eq_of_le hCon with h|h
· tauto
· tauto
let A := jensen h₂r f (h₁f.restrict (-r)) h₂f h₃f
simp at A
rw [A]
clear A
simp
have {A B : } : -A + B = B - A := by ring
rw [this]
congr 1
congr 1
let A := integrabl_congr_negRadius (f := (fun z ↦ log (Complex.abs (f z)))) (r := r)
rw [A]
have : |r| = -r := by
rw [← abs_of_pos h₂r]
simp
rw [this]
theorem Nevanlinna_firstMain₂
{f : }
{a : }
{r : }
(h₁f : MeromorphicOn f ) :
|(h₁f.T_infty r) - ((h₁f.sub (MeromorphicOn.const a)).T_infty r)| ≤ logpos ‖a‖ + log 2 := by
-- See Lang, p. 168
sorry