nevanlinna/Nevanlinna/stronglyMeromorphic_JensenFormula.lean

375 lines
12 KiB
Plaintext
Raw Normal View History

import Mathlib.Analysis.Complex.CauchyIntegral
import Mathlib.Analysis.Analytic.IsolatedZeros
import Nevanlinna.analyticOnNhd_zeroSet
import Nevanlinna.harmonicAt_examples
import Nevanlinna.harmonicAt_meanValue
import Nevanlinna.specialFunctions_CircleIntegral_affine
import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.stronglyMeromorphicOn_eliminate
import Nevanlinna.meromorphicOn_divisor
open Real
2024-12-03 17:21:22 +01:00
theorem jensen
{R : }
(hR : 0 < R)
(f : )
2024-12-03 17:21:22 +01:00
(h₁f : StronglyMeromorphicOn f (Metric.closedBall 0 R))
(h₂f : f 0 ≠ 0) :
2024-12-03 17:21:22 +01:00
log ‖f 0‖ = -∑ᶠ s, (h₁f.meromorphicOn.divisor s) * log (R * ‖s‖⁻¹) + (2 * π)⁻¹ * ∫ (x : ) in (0)..(2 * π), log ‖f (circleMap 0 R x)‖ := by
2024-12-03 17:21:22 +01:00
have h₁U : IsConnected (Metric.closedBall (0 : ) R) := by
constructor
2024-12-03 17:21:22 +01:00
· apply Metric.nonempty_closedBall.mpr
exact le_of_lt hR
· exact (convex_closedBall (0 : ) R).isPreconnected
2024-12-03 17:21:22 +01:00
have h₂U : IsCompact (Metric.closedBall (0 : ) R) :=
isCompact_closedBall 0 R
2024-12-03 17:21:22 +01:00
have h'₂f : ∃ u : (Metric.closedBall (0 : ) R), f u ≠ 0 := by
use ⟨0, Metric.mem_closedBall_self (le_of_lt hR)⟩
have h₃f : Set.Finite (Function.support h₁f.meromorphicOn.divisor) := by
exact Divisor.finiteSupport h₂U (StronglyMeromorphicOn.meromorphicOn h₁f).divisor
2024-12-03 17:21:22 +01:00
have h₄f: Function.support (fun s ↦ (h₁f.meromorphicOn.divisor s) * log (R * ‖s‖⁻¹)) ⊆ h₃f.toFinset := by
intro x
contrapose
simp
intro hx
rw [hx]
simp
rw [finsum_eq_sum_of_support_subset _ h₄f]
obtain ⟨F, h₁F, h₂F, h₃F, h₄F⟩ := MeromorphicOn.decompose₃' h₂U h₁U h₁f h'₂f
have h₁F : Function.mulSupport (fun u ↦ fun z => (z - u) ^ (h₁f.meromorphicOn.divisor u)) ⊆ h₃f.toFinset := by
intro u
contrapose
simp
intro hu
rw [hu]
simp
exact rfl
rw [finprod_eq_prod_of_mulSupport_subset _ h₁F] at h₄F
let G := fun z ↦ log ‖F z‖ + ∑ᶠ s, (h₁f.meromorphicOn.divisor s) * log ‖z - s‖
have h₁G {z : } : Function.support (fun s ↦ (h₁f.meromorphicOn.divisor s) * log ‖z - s‖) ⊆ h₃f.toFinset := by
intro s
contrapose
simp
intro hs
rw [hs]
simp
2024-12-03 17:21:22 +01:00
have decompose_f : ∀ z ∈ Metric.closedBall (0 : ) R, f z ≠ 0 → log ‖f z‖ = G z := by
intro z h₁z h₂z
rw [h₄F]
simp only [Pi.mul_apply, norm_mul]
simp only [Finset.prod_apply, norm_prod, norm_zpow]
rw [Real.log_mul]
rw [Real.log_prod]
simp_rw [Real.log_zpow]
dsimp only [G]
rw [finsum_eq_sum_of_support_subset _ h₁G]
--
intro x hx
have : z ≠ x := by
by_contra hCon
rw [← hCon] at hx
simp at hx
rw [← StronglyMeromorphicAt.order_eq_zero_iff] at h₂z
unfold MeromorphicOn.divisor at hx
simp [h₁z] at hx
tauto
apply zpow_ne_zero
simpa
-- Complex.abs (F z) ≠ 0
simp
exact h₃F ⟨z, h₁z⟩
--
rw [Finset.prod_ne_zero_iff]
intro x hx
have : z ≠ x := by
by_contra hCon
rw [← hCon] at hx
simp at hx
rw [← StronglyMeromorphicAt.order_eq_zero_iff] at h₂z
unfold MeromorphicOn.divisor at hx
simp [h₁z] at hx
tauto
apply zpow_ne_zero
simpa
2024-12-03 17:21:22 +01:00
have int_logAbs_f_eq_int_G : ∫ (x : ) in (0)..2 * π, log ‖f (circleMap 0 R x)‖ = ∫ (x : ) in (0)..2 * π, G (circleMap 0 R x) := by
rw [intervalIntegral.integral_congr_ae]
rw [MeasureTheory.ae_iff]
apply Set.Countable.measure_zero
simp
2024-12-03 17:21:22 +01:00
have t₀ : {a | a ∈ Ι 0 (2 * π) ∧ ¬log ‖f (circleMap 0 R a)‖ = G (circleMap 0 R a)}
⊆ (circleMap 0 R)⁻¹' (h₃f.toFinset) := by
intro a ha
simp at ha
simp
by_contra C
2024-12-03 17:21:22 +01:00
have t₀ : (circleMap 0 R a) ∈ Metric.closedBall 0 R := by
apply circleMap_mem_closedBall
exact le_of_lt hR
have t₁ : f (circleMap 0 R a) ≠ 0 := by
let A := h₁f (circleMap 0 R a) t₀
rw [← A.order_eq_zero_iff]
unfold MeromorphicOn.divisor at C
simp [t₀] at C
rcases C with C₁|C₂
· assumption
· let B := h₁f.meromorphicOn.order_ne_top' h₁U
2024-12-03 17:21:22 +01:00
let C := fun q ↦ B q ⟨(circleMap 0 R a), t₀⟩
rw [C₂] at C
2024-12-03 17:21:22 +01:00
have : ∃ u : (Metric.closedBall (0 : ) R), (h₁f u u.2).meromorphicAt.order ≠ := by
use ⟨(0 : ), (by simp; exact le_of_lt hR)⟩
let H := h₁f 0 (by simp; exact le_of_lt hR)
let K := H.order_eq_zero_iff.2 h₂f
rw [K]
simp
let D := C this
tauto
2024-12-03 17:21:22 +01:00
exact ha.2 (decompose_f (circleMap 0 R a) t₀ t₁)
apply Set.Countable.mono t₀
apply Set.Countable.preimage_circleMap
apply Set.Finite.countable
exact Finset.finite_toSet h₃f.toFinset
--
simp
2024-12-03 17:21:22 +01:00
have decompose_int_G : ∫ (x : ) in (0)..2 * π, G (circleMap 0 R x)
= (∫ (x : ) in (0)..2 * π, log (Complex.abs (F (circleMap 0 R x))))
+ ∑ᶠ x, (h₁f.meromorphicOn.divisor x) * ∫ (x_1 : ) in (0)..2 * π, log (Complex.abs (circleMap 0 R x_1 - ↑x)) := by
dsimp [G]
rw [intervalIntegral.integral_add]
congr
2024-12-03 17:21:22 +01:00
have t₀ {x : } : Function.support (fun s ↦ (h₁f.meromorphicOn.divisor s) * log (Complex.abs (circleMap 0 R x - s))) ⊆ h₃f.toFinset := by
intro s hs
simp at hs
simp [hs.1]
conv =>
left
arg 1
intro x
rw [finsum_eq_sum_of_support_subset _ t₀]
rw [intervalIntegral.integral_finset_sum]
2024-12-03 17:21:22 +01:00
let G' := fun x ↦ ((h₁f.meromorphicOn.divisor x) : ) * ∫ (x_1 : ) in (0)..2 * π, log (Complex.abs (circleMap 0 R x_1 - x))
have t₁ : (Function.support fun x ↦ (h₁f.meromorphicOn.divisor x) * ∫ (x_1 : ) in (0)..2 * π, log (Complex.abs (circleMap 0 R x_1 - x))) ⊆ h₃f.toFinset := by
simp
intro s
contrapose!
simp
tauto
conv =>
right
rw [finsum_eq_sum_of_support_subset _ t₁]
simp
-- ∀ i ∈ (finiteZeros h₁U h₂U h'₁f h'₂f).toFinset,
-- IntervalIntegrable (fun x => (h'₁f.order i).toNat *
-- log (Complex.abs (circleMap 0 1 x - ↑i))) MeasureTheory.volume 0 (2 * π)
intro i _
apply IntervalIntegrable.const_mul
--simp at this
2024-12-03 17:21:22 +01:00
by_cases h₂i : ‖i‖ = R
-- case pos
2024-12-03 17:21:22 +01:00
sorry
--exact int'₂ h₂i
-- case neg
apply Continuous.intervalIntegrable
apply continuous_iff_continuousAt.2
intro x
2024-12-03 17:21:22 +01:00
have : (fun x => log (Complex.abs (circleMap 0 R x - ↑i))) = log ∘ Complex.abs ∘ (fun x ↦ circleMap 0 R x - ↑i) :=
rfl
rw [this]
apply ContinuousAt.comp
apply Real.continuousAt_log
simp
by_contra ha'
conv at h₂i =>
arg 1
rw [← ha']
rw [Complex.norm_eq_abs]
2024-12-03 17:21:22 +01:00
rw [abs_circleMap_zero R x]
simp
2024-12-03 17:21:22 +01:00
linarith
--
apply ContinuousAt.comp
apply Complex.continuous_abs.continuousAt
fun_prop
-- IntervalIntegrable (fun x => log (Complex.abs (F (circleMap 0 1 x)))) MeasureTheory.volume 0 (2 * π)
apply Continuous.intervalIntegrable
apply continuous_iff_continuousAt.2
intro x
2024-12-03 17:21:22 +01:00
have : (fun x => log (Complex.abs (F (circleMap 0 R x)))) = log ∘ Complex.abs ∘ F ∘ (fun x ↦ circleMap 0 R x) :=
rfl
rw [this]
apply ContinuousAt.comp
apply Real.continuousAt_log
simp
2024-12-03 17:21:22 +01:00
have : (circleMap 0 R x) ∈ (Metric.closedBall 0 R) := by
simp
rw [abs_le]
simp [hR]
exact le_of_lt hR
exact h₃F ⟨(circleMap 0 R x), this⟩
-- ContinuousAt (⇑Complex.abs ∘ F ∘ fun x => circleMap 0 1 x) x
apply ContinuousAt.comp
apply Complex.continuous_abs.continuousAt
apply ContinuousAt.comp
2024-12-03 16:54:13 +01:00
apply DifferentiableAt.continuousAt (𝕜 := )
apply AnalyticAt.differentiableAt
2024-12-03 17:21:22 +01:00
apply h₂F (circleMap 0 R x)
simp; rw [abs_le]; simp [hR]; exact le_of_lt hR
-- ContinuousAt (fun x => circleMap 0 1 x) x
apply Continuous.continuousAt
apply continuous_circleMap
2024-12-03 16:54:13 +01:00
-- IntervalIntegrable (fun x => ∑ᶠ (s : ), ↑(↑⋯.divisor s) * log (Complex.abs (circleMap 0 1 x - s))) MeasureTheory.volume 0 (2 * π)
--simp? at h₁G
have h₁G' {z : } : (Function.support fun s => (h₁f.meromorphicOn.divisor s) * log (Complex.abs (z - s))) ⊆ ↑h₃f.toFinset := by
exact h₁G
conv =>
arg 1
intro z
rw [finsum_eq_sum_of_support_subset _ h₁G']
conv =>
arg 1
rw [← Finset.sum_fn]
apply IntervalIntegrable.sum
intro i _
apply IntervalIntegrable.const_mul
--have : i.1 ∈ Metric.closedBall (0 : ) 1 := i.2
--simp at this
2024-12-03 17:21:22 +01:00
by_cases h₂i : ‖i‖ = R
-- case pos
2024-12-03 17:21:22 +01:00
--exact int'₂ h₂i
sorry
-- case neg
--have : i.1 ∈ Metric.ball (0 : ) 1 := by sorry
apply Continuous.intervalIntegrable
apply continuous_iff_continuousAt.2
intro x
2024-12-03 17:21:22 +01:00
have : (fun x => log (Complex.abs (circleMap 0 R x - ↑i))) = log ∘ Complex.abs ∘ (fun x ↦ circleMap 0 R x - ↑i) :=
rfl
rw [this]
apply ContinuousAt.comp
apply Real.continuousAt_log
simp
by_contra ha'
conv at h₂i =>
arg 1
rw [← ha']
rw [Complex.norm_eq_abs]
2024-12-03 17:21:22 +01:00
rw [abs_circleMap_zero R x]
simp
2024-12-03 17:21:22 +01:00
linarith
apply ContinuousAt.comp
apply Complex.continuous_abs.continuousAt
fun_prop
2024-12-03 17:21:22 +01:00
have t₁ : (∫ (x : ) in (0)..2 * Real.pi, log ‖F (circleMap 0 R x)‖) = 2 * Real.pi * log ‖F 0‖ := by
let logAbsF := fun w ↦ Real.log ‖F w‖
2024-12-03 17:21:22 +01:00
have t₀ : ∀ z ∈ Metric.closedBall 0 R, HarmonicAt logAbsF z := by
intro z hz
apply logabs_of_holomorphicAt_is_harmonic
2024-12-03 16:54:13 +01:00
exact AnalyticAt.holomorphicAt (h₂F z hz)
exact h₃F ⟨z, hz⟩
2024-12-03 17:21:22 +01:00
apply harmonic_meanValue₁ R hR t₀
simp_rw [← Complex.norm_eq_abs] at decompose_int_G
rw [t₁] at decompose_int_G
2024-12-03 16:54:13 +01:00
2024-12-03 17:21:22 +01:00
have h₁G' : (Function.support fun s => (h₁f.meromorphicOn.divisor s) * ∫ (x_1 : ) in (0)..(2 * π), log ‖circleMap 0 R x_1 - s‖) ⊆ ↑h₃f.toFinset := by
2024-12-03 16:54:13 +01:00
intro s hs
simp at hs
simp [hs.1]
rw [finsum_eq_sum_of_support_subset _ h₁G'] at decompose_int_G
2024-12-03 17:21:22 +01:00
have : ∑ s ∈ h₃f.toFinset, (h₁f.meromorphicOn.divisor s) * ∫ (x_1 : ) in (0)..(2 * π), log ‖circleMap 0 R x_1 - s‖ = 0 := by
2024-12-03 16:54:13 +01:00
apply Finset.sum_eq_zero
intro x hx
rw [int₃ _]
simp
simp at hx
let ZZ := h₁f.meromorphicOn.divisor.supportInU
simp at ZZ
let UU := ZZ x hx
simpa
rw [this] at decompose_int_G
simp at decompose_int_G
rw [int_logAbs_f_eq_int_G]
rw [decompose_int_G]
2024-12-03 16:54:13 +01:00
let X := h₄F
nth_rw 1 [h₄F]
simp
have {l : } : π⁻¹ * 2⁻¹ * (2 * π * l) = l := by
calc π⁻¹ * 2⁻¹ * (2 * π * l)
_ = π⁻¹ * (2⁻¹ * 2) * π * l := by ring
_ = π⁻¹ * π * l := by ring
_ = (π⁻¹ * π) * l := by ring
_ = 1 * l := by
rw [inv_mul_cancel₀]
exact pi_ne_zero
_ = l := by simp
rw [this]
rw [log_mul]
rw [log_prod]
simp
2024-12-03 16:54:13 +01:00
rw [add_comm]
--
intro x hx
simp at hx
2024-12-03 16:54:13 +01:00
rw [zpow_ne_zero_iff]
by_contra hCon
simp at hCon
rw [← (h₁f 0 (by simp)).order_eq_zero_iff] at h₂f
rw [hCon] at hx
unfold MeromorphicOn.divisor at hx
simp at hx
rw [h₂f] at hx
tauto
2024-12-03 16:54:13 +01:00
assumption
--
simp
2024-12-03 16:54:13 +01:00
by_contra hCon
nth_rw 1 [h₄F] at h₂f
simp at h₂f
tauto
--
rw [Finset.prod_ne_zero_iff]
intro x hx
simp at hx
2024-12-03 16:54:13 +01:00
rw [zpow_ne_zero_iff]
by_contra hCon
simp at hCon
rw [← (h₁f 0 (by simp)).order_eq_zero_iff] at h₂f
rw [hCon] at hx
unfold MeromorphicOn.divisor at hx
simp at hx
rw [h₂f] at hx
tauto
2024-12-03 16:54:13 +01:00
assumption