nevanlinna/Nevanlinna/firstMain.lean

155 lines
4.0 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.stronglyMeromorphicOn
import Nevanlinna.meromorphicOn_divisor
2024-12-19 16:10:51 +01:00
import Nevanlinna.meromorphicOn_integrability
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-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-09 19:58:56 +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-09 19:58:56 +01:00
fun r ↦ ∑ᶠ z, (max 0 (-((hf.restrict r).divisor z))) * log (r * ‖z‖⁻¹)
2024-12-05 16:53:48 +01:00
theorem Nevanlinna_counting
{f : }
2024-12-09 19:58:56 +01:00
(hf : MeromorphicOn f ) :
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
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
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
theorem Nevanlinna_proximity₀
{f : }
{r : }
(h₁f : MeromorphicOn f )
2024-12-19 21:24:13 +01:00
(hr : 0 < r) :
2024-12-19 16:10:51 +01:00
(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 hr
2024-12-19 21:24:13 +01:00
intro z hx
exact h₁f z trivial
--
apply MeromorphicOn.integrable_logpos_abs_f hr
exact MeromorphicOn.inv_iff.mpr fun x a => h₁f x trivial
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
by_cases h₁r : r = 0
· unfold MeromorphicOn.m_infty
rw [← mul_sub]; congr
simp only [h₁r, circleMap_zero_radius, Function.const_apply, intervalIntegral.integral_const, sub_zero, smul_eq_mul, Pi.inv_apply, norm_inv]
rw [← mul_sub]; congr
exact loglogpos
2024-12-19 21:24:13 +01:00
by_cases h₂r : 0 < r
· exact Nevanlinna_proximity₀ h₁f h₂r
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-19 16:10:51 +01:00
2024-12-19 21:24:13 +01:00
sorry
2024-12-05 16:53:48 +01:00
sorry
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
theorem Nevanlinna_firstMain₁
{f : }
(h₁f : MeromorphicOn f )
(h₂f : StronglyMeromorphicAt f 0)
(h₃f : f 0 ≠ 0) :
(fun r ↦ log ‖f 0‖) + h₁f.inv.T_infty = h₁f.T_infty := by
funext r
simp
unfold MeromorphicOn.T_infty
unfold MeromorphicOn.N_infty
unfold MeromorphicOn.m_infty
simp
sorry
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
sorry