2024-07-31 14:54:41 +02:00
|
|
|
|
import Nevanlinna.harmonicAt_examples
|
|
|
|
|
import Nevanlinna.harmonicAt_meanValue
|
|
|
|
|
|
2024-08-01 10:02:36 +02:00
|
|
|
|
lemma int₀
|
|
|
|
|
{a : ℂ}
|
|
|
|
|
(ha : a ∈ Metric.ball 0 1)
|
|
|
|
|
:
|
|
|
|
|
∫ (x : ℝ) in (0)..2 * Real.pi, Real.log ‖circleMap 0 1 x - a‖ = 0 := by
|
|
|
|
|
sorry
|
|
|
|
|
|
2024-07-31 14:54:41 +02:00
|
|
|
|
|
|
|
|
|
theorem jensen_case_R_eq_one
|
|
|
|
|
(f : ℂ → ℂ)
|
|
|
|
|
(h₁f : Differentiable ℂ f)
|
|
|
|
|
(h₂f : f 0 ≠ 0)
|
|
|
|
|
(S : Finset ℕ)
|
|
|
|
|
(a : S → ℂ)
|
|
|
|
|
(ha : ∀ s, a s ∈ Metric.ball 0 1)
|
|
|
|
|
(F : ℂ → ℂ)
|
|
|
|
|
(h₁F : Differentiable ℂ F)
|
|
|
|
|
(h₂F : ∀ z, F z ≠ 0)
|
2024-07-31 15:59:43 +02:00
|
|
|
|
(h₃F : f = fun z ↦ (F z) * ∏ s : S, (z - a s))
|
2024-07-31 14:54:41 +02:00
|
|
|
|
:
|
|
|
|
|
Real.log ‖f 0‖ = -∑ s, Real.log (‖a s‖⁻¹) + (2 * Real.pi)⁻¹ * ∫ (x : ℝ) in (0)..2 * Real.pi, Real.log ‖f (circleMap 0 1 x)‖ := by
|
|
|
|
|
|
|
|
|
|
let logAbsF := fun w ↦ Real.log ‖F w‖
|
|
|
|
|
|
|
|
|
|
have t₀ : ∀ z, HarmonicAt logAbsF z := by
|
|
|
|
|
intro z
|
|
|
|
|
apply logabs_of_holomorphicAt_is_harmonic
|
2024-07-31 15:59:43 +02:00
|
|
|
|
apply h₁F.holomorphicAt
|
2024-07-31 14:54:41 +02:00
|
|
|
|
exact h₂F z
|
|
|
|
|
|
|
|
|
|
have t₁ : (∫ (x : ℝ) in (0)..2 * Real.pi, logAbsF (circleMap 0 1 x)) = 2 * Real.pi * logAbsF 0 := by
|
|
|
|
|
apply harmonic_meanValue t₀ 1
|
|
|
|
|
exact Real.zero_lt_one
|
|
|
|
|
|
2024-07-31 15:59:43 +02:00
|
|
|
|
have t₂ : ∀ s, f (a s) = 0 := by
|
|
|
|
|
intro s
|
|
|
|
|
rw [h₃F]
|
|
|
|
|
simp
|
|
|
|
|
right
|
|
|
|
|
apply Finset.prod_eq_zero_iff.2
|
|
|
|
|
use s
|
|
|
|
|
simp
|
|
|
|
|
|
2024-07-31 14:54:41 +02:00
|
|
|
|
let logAbsf := fun w ↦ Real.log ‖f w‖
|
|
|
|
|
have s₀ : ∀ z, f z ≠ 0 → logAbsf z = logAbsF z + ∑ s, Real.log ‖z - a s‖ := by
|
2024-07-31 15:59:43 +02:00
|
|
|
|
intro z hz
|
|
|
|
|
dsimp [logAbsf]
|
|
|
|
|
rw [h₃F]
|
|
|
|
|
simp_rw [Complex.abs.map_mul]
|
|
|
|
|
rw [Complex.abs_prod]
|
|
|
|
|
rw [Real.log_mul]
|
|
|
|
|
rw [Real.log_prod]
|
|
|
|
|
rfl
|
|
|
|
|
intro s hs
|
|
|
|
|
simp
|
|
|
|
|
by_contra ha'
|
|
|
|
|
rw [ha'] at hz
|
|
|
|
|
exact hz (t₂ s)
|
|
|
|
|
-- Complex.abs (F z) ≠ 0
|
|
|
|
|
simp
|
|
|
|
|
exact h₂F z
|
|
|
|
|
-- ∏ I : { x // x ∈ S }, Complex.abs (z - a I) ≠ 0
|
|
|
|
|
by_contra h'
|
|
|
|
|
obtain ⟨s, h's, h''⟩ := Finset.prod_eq_zero_iff.1 h'
|
|
|
|
|
simp at h''
|
|
|
|
|
rw [h''] at hz
|
|
|
|
|
let A := t₂ s
|
|
|
|
|
exact hz A
|
|
|
|
|
|
2024-07-31 14:54:41 +02:00
|
|
|
|
have s₁ : ∀ z, f z ≠ 0 → logAbsF z = logAbsf z - ∑ s, Real.log ‖z - a s‖ := by
|
2024-08-01 10:02:36 +02:00
|
|
|
|
intro z hz
|
|
|
|
|
rw [s₀ z hz]
|
|
|
|
|
simp
|
2024-07-31 14:54:41 +02:00
|
|
|
|
|
|
|
|
|
rw [s₁ 0 h₂f] at t₁
|
|
|
|
|
|
2024-08-01 10:02:36 +02:00
|
|
|
|
have h₀ {x : ℝ} : f (circleMap 0 1 x) ≠ 0 := by
|
|
|
|
|
rw [h₃F]
|
|
|
|
|
simp
|
|
|
|
|
constructor
|
|
|
|
|
· exact h₂F (circleMap 0 1 x)
|
|
|
|
|
· by_contra h'
|
|
|
|
|
obtain ⟨s, _, h₂s⟩ := Finset.prod_eq_zero_iff.1 h'
|
|
|
|
|
have : circleMap 0 1 x = a s := by
|
|
|
|
|
rw [← sub_zero (circleMap 0 1 x)]
|
|
|
|
|
nth_rw 2 [← h₂s]
|
|
|
|
|
simp
|
|
|
|
|
let A := ha s
|
|
|
|
|
rw [← this] at A
|
|
|
|
|
simp at A
|
2024-07-31 15:59:43 +02:00
|
|
|
|
|
2024-08-01 10:02:36 +02:00
|
|
|
|
simp_rw [s₁ (circleMap 0 1 _) h₀] at t₁
|
2024-07-31 14:54:41 +02:00
|
|
|
|
rw [intervalIntegral.integral_sub] at t₁
|
|
|
|
|
rw [intervalIntegral.integral_finset_sum] at t₁
|
|
|
|
|
|
2024-08-01 10:02:36 +02:00
|
|
|
|
simp_rw [int₀ (ha _)] at t₁
|
2024-07-31 14:54:41 +02:00
|
|
|
|
simp at t₁
|
|
|
|
|
rw [t₁]
|
|
|
|
|
simp
|
2024-07-31 15:08:27 +02:00
|
|
|
|
have {w : ℝ} : Real.pi⁻¹ * 2⁻¹ * (2 * Real.pi * w) = w := by
|
|
|
|
|
ring_nf
|
|
|
|
|
simp [mul_inv_cancel Real.pi_ne_zero]
|
2024-07-31 14:54:41 +02:00
|
|
|
|
rw [this]
|
|
|
|
|
simp
|
|
|
|
|
rfl
|
|
|
|
|
-- ∀ i ∈ Finset.univ, IntervalIntegrable (fun x => Real.log ‖circleMap 0 1 x - a i‖) MeasureTheory.volume 0 (2 * Real.pi)
|
2024-08-01 10:02:36 +02:00
|
|
|
|
intro i _
|
2024-07-31 14:54:41 +02:00
|
|
|
|
apply Continuous.intervalIntegrable
|
|
|
|
|
apply continuous_iff_continuousAt.2
|
|
|
|
|
intro x
|
|
|
|
|
have : (fun x => Real.log ‖circleMap 0 1 x - a i‖) = Real.log ∘ Complex.abs ∘ (fun x ↦ circleMap 0 1 x - a i) :=
|
|
|
|
|
rfl
|
|
|
|
|
rw [this]
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply Real.continuousAt_log
|
|
|
|
|
simp
|
|
|
|
|
by_contra ha'
|
|
|
|
|
let A := ha i
|
|
|
|
|
rw [← ha'] at A
|
|
|
|
|
simp at A
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply Complex.continuous_abs.continuousAt
|
|
|
|
|
fun_prop
|
|
|
|
|
-- IntervalIntegrable (fun x => logAbsf (circleMap 0 1 x)) MeasureTheory.volume 0 (2 * Real.pi)
|
|
|
|
|
apply Continuous.intervalIntegrable
|
|
|
|
|
apply continuous_iff_continuousAt.2
|
|
|
|
|
intro x
|
|
|
|
|
have : (fun x => logAbsf (circleMap 0 1 x)) = Real.log ∘ Complex.abs ∘ f ∘ (fun x ↦ circleMap 0 1 x) :=
|
|
|
|
|
rfl
|
|
|
|
|
rw [this]
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
simp
|
2024-08-01 10:02:36 +02:00
|
|
|
|
exact h₀
|
2024-07-31 14:54:41 +02:00
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply Complex.continuous_abs.continuousAt
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply h₁f.continuous.continuousAt
|
|
|
|
|
let A := continuous_circleMap 0 1
|
|
|
|
|
apply A.continuousAt
|
|
|
|
|
-- IntervalIntegrable (fun x => ∑ s : { x // x ∈ S }, Real.log ‖circleMap 0 1 x - a s‖) MeasureTheory.volume 0 (2 * Real.pi)
|
|
|
|
|
apply Continuous.intervalIntegrable
|
|
|
|
|
apply continuous_finset_sum
|
2024-08-01 10:02:36 +02:00
|
|
|
|
intro i _
|
2024-07-31 14:54:41 +02:00
|
|
|
|
apply continuous_iff_continuousAt.2
|
|
|
|
|
intro x
|
|
|
|
|
have : (fun x => Real.log ‖circleMap 0 1 x - a i‖) = Real.log ∘ Complex.abs ∘ (fun x ↦ circleMap 0 1 x - a i) :=
|
|
|
|
|
rfl
|
|
|
|
|
rw [this]
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply Real.continuousAt_log
|
|
|
|
|
simp
|
|
|
|
|
by_contra ha'
|
|
|
|
|
let A := ha i
|
|
|
|
|
rw [← ha'] at A
|
|
|
|
|
simp at A
|
|
|
|
|
apply ContinuousAt.comp
|
|
|
|
|
apply Complex.continuous_abs.continuousAt
|
|
|
|
|
fun_prop
|