nevanlinna/Nevanlinna/firstMain.lean

290 lines
6.5 KiB
Plaintext
Raw Normal View History

2024-12-05 16:53:48 +01:00
import Mathlib.MeasureTheory.Integral.CircleIntegral
2024-12-05 13:51:00 +01:00
import Nevanlinna.divisor
2024-12-05 16:53:48 +01:00
import Nevanlinna.meromorphicOn_divisor
2024-12-19 16:10:51 +01:00
import Nevanlinna.meromorphicOn_integrability
2024-12-20 09:42:56 +01:00
import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.stronglyMeromorphic_JensenFormula
2024-09-13 09:21:57 +02:00
open Real
2024-12-05 16:53:48 +01:00
-- Lang p. 164
2024-12-09 19:58:56 +01:00
theorem MeromorphicOn.restrict
{f : }
(h₁f : MeromorphicOn f )
(r : ) :
MeromorphicOn f (Metric.closedBall 0 r) := by
exact fun x a => h₁f x trivial
2024-12-20 08:16:22 +01:00
theorem MeromorphicOn.restrict_inv
{f : }
(h₁f : MeromorphicOn f )
(r : ) :
h₁f.inv.restrict r = (h₁f.restrict r).inv := by
funext x
simp
2024-12-05 16:53:48 +01:00
noncomputable def MeromorphicOn.N_zero
{f : }
2024-12-09 19:58:56 +01:00
(hf : MeromorphicOn f ) :
2024-12-05 16:53:48 +01:00
:=
2024-12-20 11:30:06 +01:00
fun r ↦ ∑ᶠ z, (max 0 ((hf.restrict |r|).divisor z)) * log (r * ‖z‖⁻¹)
2024-12-05 16:53:48 +01:00
noncomputable def MeromorphicOn.N_infty
{f : }
2024-12-09 19:58:56 +01:00
(hf : MeromorphicOn f ) :
2024-12-05 16:53:48 +01:00
:=
2024-12-20 11:30:06 +01:00
fun r ↦ ∑ᶠ z, (max 0 (-((hf.restrict |r|).divisor z))) * log (r * ‖z‖⁻¹)
2024-12-05 16:53:48 +01:00
2024-12-21 07:04:13 +01:00
2024-12-24 06:50:23 +01:00
theorem Nevanlinna_counting₁₁
{f : }
{a : }
(hf : MeromorphicOn f ) :
(hf.add (MeromorphicOn.const a)).N_infty = hf.N_infty := by
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
2025-01-03 10:38:06 +01:00
intro x hx; simp at hx
2024-12-24 06:50:23 +01:00
congr 2
2025-01-03 10:38:06 +01:00
by_cases h : 0 ≤ (hf.restrict |r|).divisor x
· simp [h]
let A := (hf.restrict |r|).divisor_add_const₁ a h
exact A
· simp at h
have h' : 0 ≤ -((hf.restrict |r|).divisor x) := by
apply Int.le_neg_of_le_neg
simp
exact Int.le_of_lt h
simp [h']
clear h'
2025-01-03 11:32:28 +01:00
have A := (hf.restrict |r|).divisor_add_const₂ a h
have A' : 0 ≤ -((MeromorphicOn.add (MeromorphicOn.restrict hf |r|) (MeromorphicOn.const a)).divisor x) := by
apply Int.le_neg_of_le_neg
simp
exact Int.le_of_lt A
simp [A']
clear A A'
exact (hf.restrict |r|).divisor_add_const₃ a h
--
intro x
contrapose
simp
intro hx
rw [hx]
tauto
--
intro x
contrapose
simp
intro hx
have : 0 ≤ (hf.restrict |r|).divisor x := by
rw [hx]
have G := (hf.restrict |r|).divisor_add_const₁ a this
clear this
simp [G]
2024-12-24 06:50:23 +01:00
2024-12-20 08:16:22 +01:00
theorem Nevanlinna_counting₀
{f : }
(hf : MeromorphicOn f ) :
hf.inv.N_infty = hf.N_zero := by
2024-12-20 11:30:06 +01:00
2024-12-20 08:16:22 +01:00
funext r
unfold MeromorphicOn.N_zero MeromorphicOn.N_infty
2024-12-20 11:30:06 +01:00
let A := (hf.restrict |r|).divisor.finiteSupport (isCompact_closedBall 0 |r|)
2024-12-20 08:16:22 +01:00
repeat
rw [finsum_eq_sum_of_support_subset (s := A.toFinset)]
apply Finset.sum_congr rfl
intro x hx
congr
2024-12-20 11:30:06 +01:00
let B := hf.restrict_inv |r|
2024-12-20 08:16:22 +01:00
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
2024-12-20 11:30:06 +01:00
rw [MeromorphicOn.divisor_inv (hf.restrict |r|)] at h₁x
2024-12-20 08:16:22 +01:00
simp at h₁x
rw [hx] at h₁x
tauto
2024-12-05 16:53:48 +01:00
theorem Nevanlinna_counting
{f : }
2024-12-09 19:58:56 +01:00
(hf : MeromorphicOn f ) :
2024-12-20 11:30:06 +01:00
hf.N_zero - hf.N_infty = fun r ↦ ∑ᶠ z, ((hf.restrict |r|).divisor z) * log (r * ‖z‖⁻¹) := by
2024-12-09 16:00:26 +01:00
funext r
simp only [Pi.sub_apply]
2024-12-09 19:58:56 +01:00
unfold MeromorphicOn.N_zero MeromorphicOn.N_infty
2024-12-20 11:30:06 +01:00
let A := (hf.restrict |r|).divisor.finiteSupport (isCompact_closedBall 0 |r|)
2024-12-09 19:58:56 +01:00
repeat
rw [finsum_eq_sum_of_support_subset (s := A.toFinset)]
rw [← Finset.sum_sub_distrib]
simp_rw [← sub_mul]
congr
funext x
congr
2024-12-20 11:30:06 +01:00
by_cases h : 0 ≤ (hf.restrict |r|).divisor x
2024-12-09 19:58:56 +01:00
· simp [h]
2024-12-20 11:30:06 +01:00
· have h' : 0 ≤ -((hf.restrict |r|).divisor x) := by
2024-12-09 19:58:56 +01:00
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
2024-12-05 16:53:48 +01:00
2024-09-14 08:38:04 +02:00
2024-12-05 16:53:48 +01:00
--
noncomputable def MeromorphicOn.m_infty
{f : }
2024-12-09 19:58:56 +01:00
(_ : MeromorphicOn f ) :
2024-12-05 16:53:48 +01:00
:=
fun r ↦ (2 * π)⁻¹ * ∫ x in (0)..(2 * π), logpos ‖f (circleMap 0 r x)‖
2024-12-19 16:10:51 +01:00
2024-12-05 16:53:48 +01:00
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
2024-12-19 16:10:51 +01:00
2024-12-05 16:53:48 +01:00
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)))
--
2024-12-19 21:24:13 +01:00
apply MeromorphicOn.integrable_logpos_abs_f
2024-12-20 08:16:22 +01:00
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
2024-12-19 16:10:51 +01:00
2024-12-19 21:24:13 +01:00
2024-12-05 16:53:48 +01:00
noncomputable def MeromorphicOn.T_infty
{f : }
(hf : MeromorphicOn f ) :
2024-12-05 13:51:00 +01:00
:=
2024-12-05 16:53:48 +01:00
hf.m_infty + hf.N_infty
2024-12-20 08:16:22 +01:00
2024-12-05 16:53:48 +01:00
theorem Nevanlinna_firstMain₁
{f : }
(h₁f : MeromorphicOn f )
(h₂f : StronglyMeromorphicAt f 0)
(h₃f : f 0 ≠ 0) :
2024-12-20 11:30:06 +01:00
(fun _ ↦ log ‖f 0‖) + h₁f.inv.T_infty = h₁f.T_infty := by
2024-12-20 08:16:22 +01:00
rw [add_eq_of_eq_sub]
unfold MeromorphicOn.T_infty
2024-12-20 09:42:56 +01:00
have {A B C D : } : A + B - (C + D) = A - C - (D - B) := by
2024-12-20 08:16:22 +01:00
ring
rw [this]
clear this
2024-12-20 09:42:56 +01:00
rw [Nevanlinna_counting₀ h₁f]
rw [Nevanlinna_counting h₁f]
2024-12-05 16:53:48 +01:00
funext r
simp
2024-12-20 08:16:22 +01:00
rw [← Nevanlinna_proximity h₁f]
2024-12-20 11:30:06 +01:00
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]
2024-12-20 08:16:22 +01:00
2024-12-05 16:53:48 +01:00
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
2024-12-20 11:30:06 +01:00
2024-12-20 11:54:02 +01:00
-- See Lang, p. 168
2024-12-05 16:53:48 +01:00
sorry