2024-07-31 12:25:30 +02:00
|
|
|
|
import Mathlib.Analysis.Complex.CauchyIntegral
|
|
|
|
|
import Nevanlinna.holomorphic_examples
|
|
|
|
|
|
|
|
|
|
theorem harmonic_meanValue
|
|
|
|
|
{f : ℂ → ℝ}
|
2024-08-08 14:26:53 +02:00
|
|
|
|
{z : ℂ}
|
|
|
|
|
(ρ R : ℝ)
|
|
|
|
|
(hR : R > 0)
|
|
|
|
|
(hρ : ρ > R)
|
|
|
|
|
(hf : ∀ x ∈ Metric.ball z ρ , HarmonicAt f x)
|
|
|
|
|
:
|
|
|
|
|
(∫ (x : ℝ) in (0)..2 * Real.pi, f (circleMap z R x)) = 2 * Real.pi * f z
|
2024-07-31 12:25:30 +02:00
|
|
|
|
:= by
|
|
|
|
|
|
2024-08-08 14:26:53 +02:00
|
|
|
|
obtain ⟨F, h₁F, h₂F⟩ := harmonic_is_realOfHolomorphic (gt_trans hρ hR) hf
|
2024-07-31 12:25:30 +02:00
|
|
|
|
|
2024-08-08 14:26:53 +02:00
|
|
|
|
have hrρ : Metric.ball z R ⊆ Metric.ball z ρ := by
|
|
|
|
|
intro x hx
|
|
|
|
|
exact gt_trans hρ hx
|
2024-07-31 12:25:30 +02:00
|
|
|
|
|
2024-08-08 14:26:53 +02:00
|
|
|
|
have reg₀F : DifferentiableOn ℂ F (Metric.ball z ρ) := by
|
|
|
|
|
intro x hx
|
|
|
|
|
apply DifferentiableAt.differentiableWithinAt
|
|
|
|
|
apply HolomorphicAt.differentiableAt (h₁F x _)
|
|
|
|
|
exact hx
|
|
|
|
|
|
|
|
|
|
have reg₁F : DifferentiableOn ℂ F (Metric.ball z R) := by
|
|
|
|
|
intro x hx
|
|
|
|
|
apply DifferentiableAt.differentiableWithinAt
|
|
|
|
|
apply HolomorphicAt.differentiableAt (h₁F x _)
|
|
|
|
|
exact hrρ hx
|
|
|
|
|
|
|
|
|
|
have : (∮ (x : ℂ) in C(z, R), (x - z)⁻¹ • F x) = (2 * ↑Real.pi * Complex.I) • F z := by
|
2024-07-31 12:25:30 +02:00
|
|
|
|
let s : Set ℂ := ∅
|
|
|
|
|
let hs : s.Countable := Set.countable_empty
|
|
|
|
|
let _ : ℂ := 0
|
2024-08-08 14:26:53 +02:00
|
|
|
|
have hw : (z : ℂ) ∈ Metric.ball z R := Metric.mem_ball_self hR
|
|
|
|
|
have hc : ContinuousOn F (Metric.closedBall z R) := by
|
|
|
|
|
apply reg₀F.continuousOn.mono
|
|
|
|
|
intro x hx
|
|
|
|
|
simp at hx
|
|
|
|
|
simp
|
|
|
|
|
linarith
|
|
|
|
|
have hd : ∀ x ∈ Metric.ball z R \ s, DifferentiableAt ℂ F x := by
|
|
|
|
|
intro x hx
|
|
|
|
|
let A := reg₁F x hx.1
|
|
|
|
|
apply A.differentiableAt
|
|
|
|
|
apply (IsOpen.mem_nhds_iff ?hs).mpr
|
|
|
|
|
exact hx.1
|
|
|
|
|
exact Metric.isOpen_ball
|
2024-07-31 12:25:30 +02:00
|
|
|
|
let CIF := Complex.circleIntegral_sub_inv_smul_of_differentiable_on_off_countable hs hw hc hd
|
|
|
|
|
simp at CIF
|
|
|
|
|
assumption
|
|
|
|
|
|
|
|
|
|
unfold circleIntegral at this
|
|
|
|
|
simp_rw [deriv_circleMap] at this
|
|
|
|
|
|
2024-08-08 14:26:53 +02:00
|
|
|
|
have t₁ {θ : ℝ} : (circleMap 0 R θ * Complex.I) • (circleMap 0 R θ)⁻¹ • F (circleMap z R θ) = Complex.I • F (circleMap z R θ) := by
|
2024-07-31 12:25:30 +02:00
|
|
|
|
rw [← smul_assoc]
|
|
|
|
|
congr 1
|
|
|
|
|
simp
|
|
|
|
|
nth_rw 1 [mul_comm]
|
|
|
|
|
rw [← mul_assoc]
|
|
|
|
|
simp
|
|
|
|
|
apply inv_mul_cancel
|
|
|
|
|
apply circleMap_ne_center
|
|
|
|
|
exact Ne.symm (ne_of_lt hR)
|
|
|
|
|
simp_rw [t₁] at this
|
|
|
|
|
simp at this
|
|
|
|
|
|
|
|
|
|
have t₂ : Complex.reCLM (-Complex.I * (Complex.I * ∫ (x : ℝ) in (0)..2 * Real.pi, F (circleMap 0 R x))) = Complex.reCLM (-Complex.I * (2 * ↑Real.pi * Complex.I * F 0)) := by
|
|
|
|
|
rw [this]
|
|
|
|
|
simp at t₂
|
|
|
|
|
have xx {z : ℂ} : (F z).re = f z := by
|
|
|
|
|
rw [← h₂F]
|
|
|
|
|
simp
|
|
|
|
|
simp_rw [xx] at t₂
|
|
|
|
|
have x₁ {z : ℂ} : z.re = Complex.reCLM z := by rfl
|
|
|
|
|
rw [x₁] at t₂
|
|
|
|
|
rw [← ContinuousLinearMap.intervalIntegral_comp_comm] at t₂
|
|
|
|
|
simp at t₂
|
|
|
|
|
simp_rw [xx] at t₂
|
|
|
|
|
exact t₂
|
|
|
|
|
-- IntervalIntegrable (fun x => F (circleMap 0 1 x)) MeasureTheory.volume 0 (2 * Real.pi)
|
|
|
|
|
apply Continuous.intervalIntegrable
|
|
|
|
|
apply Continuous.comp
|
|
|
|
|
exact regF.continuous
|
|
|
|
|
exact continuous_circleMap 0 R
|