import Mathlib.Analysis.Complex.CauchyIntegral import Nevanlinna.holomorphic_examples theorem harmonic_meanValue {f : ℂ → ℝ} (hf : ∀ z, HarmonicAt f z) (R : ℝ) (hR : R > 0) : (∫ (x : ℝ) in (0)..2 * Real.pi, f (circleMap 0 R x)) = 2 * Real.pi * f 0 := by obtain ⟨F, h₁F, h₂F⟩ := harmonic_is_realOfHolomorphic hf have regF : Differentiable ℂ F := fun z ↦ HolomorphicAt.differentiableAt (h₁F z) have : (∮ (z : ℂ) in C(0, R), z⁻¹ • F z) = (2 * ↑Real.pi * Complex.I) • F 0 := by let s : Set ℂ := ∅ let hs : s.Countable := Set.countable_empty let _ : ℂ := 0 let hw : (0 : ℂ) ∈ Metric.ball 0 R := Metric.mem_ball_self hR let hc : ContinuousOn F (Metric.closedBall 0 R) := by apply Continuous.continuousOn exact regF.continuous let hd : ∀ x ∈ Metric.ball 0 R \ s, DifferentiableAt ℂ F x := by intro x _ exact regF x 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 have t₁ {θ : ℝ} : (circleMap 0 R θ * Complex.I) • (circleMap 0 R θ)⁻¹ • F (circleMap 0 R θ) = Complex.I • F (circleMap 0 R θ) := by 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