Update harmonicAt_meanValue.lean

This commit is contained in:
Stefan Kebekus 2024-08-08 15:04:41 +02:00
parent 4642c017c7
commit 4b6cdcc76a
1 changed files with 27 additions and 8 deletions

View File

@ -65,24 +65,43 @@ theorem harmonic_meanValue
apply inv_mul_cancel
apply circleMap_ne_center
exact Ne.symm (ne_of_lt hR)
simp_rw [t₁] at this
have t'₁ {θ : } : circleMap 0 R θ = circleMap z R θ - z := by
exact Eq.symm (circleMap_sub_center z R θ)
simp_rw [← t'₁, 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
have t₂ : Complex.reCLM (-Complex.I * (Complex.I * ∫ (x : ) in (0)..2 * Real.pi, F (circleMap z R x))) = Complex.reCLM (-Complex.I * (2 * ↑Real.pi * Complex.I * F z)) := by
rw [this]
simp at t₂
have xx {z : } : (F z).re = f z := by
have xx {x : } : (F (circleMap z R x)).re = f (circleMap z R x) := by
rw [← h₂F]
simp
simp_rw [xx] at t₂
simp
rw [Complex.dist_eq]
rw [circleMap_sub_center]
simp
rwa [abs_of_nonneg (le_of_lt hR)]
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₂
have x₁ {z : } : z.re = Complex.reCLM z := by rfl
rw [x₁] at t₂
have : Complex.reCLM (F z) = f z := by
apply h₂F
simp
exact gt_trans hρ hR
rw [this] 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
apply ContinuousOn.intervalIntegrable
apply ContinuousOn.comp reg₀F.continuousOn _
intro θ _
simp
rw [Complex.dist_eq, circleMap_sub_center]
simp
rwa [abs_of_nonneg (le_of_lt hR)]
apply Continuous.continuousOn
exact continuous_circleMap z R