This commit is contained in:
Stefan Kebekus 2024-08-08 17:45:07 +02:00
parent 4b6cdcc76a
commit 17705601c2
2 changed files with 45 additions and 8 deletions

View File

@ -5,8 +5,8 @@ theorem harmonic_meanValue
{f : } {f : }
{z : } {z : }
(ρ R : ) (ρ R : )
(hR : R > 0) (hR : 0 < R)
(hρ : ρ > R) (hρ : R < ρ)
(hf : ∀ x ∈ Metric.ball z ρ , HarmonicAt f x) (hf : ∀ x ∈ Metric.ball z ρ , HarmonicAt f x)
: :
(∫ (x : ) in (0)..2 * Real.pi, f (circleMap z R x)) = 2 * Real.pi * f z (∫ (x : ) in (0)..2 * Real.pi, f (circleMap z R x)) = 2 * Real.pi * f z

View File

@ -1,13 +1,48 @@
import Nevanlinna.harmonicAt_examples import Nevanlinna.harmonicAt_examples
import Nevanlinna.harmonicAt_meanValue import Nevanlinna.harmonicAt_meanValue
lemma l₀
{x₁ x₂ : } :
(circleMap 0 1 x₁) * (circleMap 0 1 x₂) = circleMap 0 1 (x₁+x₂) := by
dsimp [circleMap]
simp
rw [add_mul, Complex.exp_add]
lemma int₀ lemma int₀
{a : } {a : }
(ha : a ∈ Metric.ball 0 1) (ha : a ∈ Metric.ball 0 1) :
:
∫ (x : ) in (0)..2 * Real.pi, Real.log ‖circleMap 0 1 x - a‖ = 0 := by ∫ (x : ) in (0)..2 * Real.pi, Real.log ‖circleMap 0 1 x - a‖ = 0 := by
have {x : } : ‖(circleMap 0 1 x) - a‖ = ‖(circleMap 0 1 x) - a‖ := by
calc ‖(circleMap 0 1 x) - a‖
_ = 1 * ‖(circleMap 0 1 x) - a‖ := by exact Eq.symm (one_mul ‖circleMap 0 1 x - a‖)
_ = ‖(circleMap 0 1 (-x))‖ * ‖(circleMap 0 1 x) - a‖ := by
have : ‖(circleMap 0 1 (-x))‖ = 1 := by
rw [Complex.norm_eq_abs, abs_circleMap_zero]
simp
rw [this]
_ = ‖(circleMap 0 1 (-x)) * ((circleMap 0 1 x) - a)‖ := by
exact Eq.symm (NormedField.norm_mul' (circleMap 0 1 (-x)) (circleMap 0 1 x - a))
_ = ‖(circleMap 0 1 (-x)) * (circleMap 0 1 x) - (circleMap 0 1 (-x)) * a‖ := by
rw [mul_sub]
_ =
sorry sorry
conv =>
left
arg 1
intro x
rw [← this]
simp
have hf : ∀ x ∈ Metric.ball 0 2, HarmonicAt F x := by sorry
sorry
theorem jensen_case_R_eq_one theorem jensen_case_R_eq_one
(f : ) (f : )
@ -25,15 +60,17 @@ theorem jensen_case_R_eq_one
let logAbsF := fun w ↦ Real.log ‖F w‖ let logAbsF := fun w ↦ Real.log ‖F w‖
have t₀ : ∀ z, HarmonicAt logAbsF z := by have t₀ : ∀ z ∈ Metric.ball 0 2, HarmonicAt logAbsF z := by
intro z intro z _
apply logabs_of_holomorphicAt_is_harmonic apply logabs_of_holomorphicAt_is_harmonic
apply h₁F.holomorphicAt apply h₁F.holomorphicAt
exact h₂F z exact h₂F z
have t₁ : (∫ (x : ) in (0)..2 * Real.pi, logAbsF (circleMap 0 1 x)) = 2 * Real.pi * logAbsF 0 := by have t₁ : (∫ (x : ) in (0)..2 * Real.pi, logAbsF (circleMap 0 1 x)) = 2 * Real.pi * logAbsF 0 := by
apply harmonic_meanValue t₀ 1 have hR : (0 : ) < (1 : ) := by apply Real.zero_lt_one
exact Real.zero_lt_one have hρ : (1 : ) < (2 : ) := by linarith
apply harmonic_meanValue 2 1 hR hρ t₀
have t₂ : ∀ s, f (a s) = 0 := by have t₂ : ∀ s, f (a s) = 0 := by
intro s intro s