Compare commits

...

2 Commits

Author SHA1 Message Date
Stefan Kebekus ac3cd65bf2 Update complexHarmonic.lean 2024-05-30 10:49:16 +02:00
Stefan Kebekus 7f10e28525 Update complexHarmonic.lean 2024-05-30 10:41:10 +02:00
1 changed files with 15 additions and 14 deletions

View File

@ -89,7 +89,7 @@ theorem harmonic_comp_CLM_is_harmonic {f : → F₁} {l : F₁ →L[] G}
exact ContDiff.restrict_scalars h.1 exact ContDiff.restrict_scalars h.1
theorem harmonicOn_comp_CLM_is_harmonicOn {f : → F₁} {s : Set } {l : F₁ →L[] G} (h : HarmonicOn f s) : theorem harmonicOn_comp_CLM_is_harmonicOn {f : → F₁} {s : Set } {l : F₁ →L[] G} (hs : IsOpen s) (h : HarmonicOn f s) :
HarmonicOn (l ∘ f) s := by HarmonicOn (l ∘ f) s := by
constructor constructor
@ -97,15 +97,14 @@ theorem harmonicOn_comp_CLM_is_harmonicOn {f : → F₁} {s : Set } {l :
apply ContDiffOn.continuousLinearMap_comp apply ContDiffOn.continuousLinearMap_comp
exact h.1 exact h.1
· -- Vanishing of Laplace · -- Vanishing of Laplace
rw [laplace_compContLin]
simp
intro z zHyp intro z zHyp
rw [laplace_compContLinAt]
simp
rw [h.2 z] rw [h.2 z]
simp simp
assumption assumption
apply ContDiffOn.contDiffAt h.1
exact IsOpen.mem_nhds hs zHyp
theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[] G₁} : theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[] G₁} :
@ -200,15 +199,16 @@ theorem antiholomorphic_is_harmonic {f : } (h : Differentiable f)
theorem log_normSq_of_holomorphicOn_is_harmonicOn theorem log_normSq_of_holomorphicOn_is_harmonicOn
{f : } {f : }
{s : Set } {s : Set }
(hs : IsOpen s)
(h₁ : DifferentiableOn f s) (h₁ : DifferentiableOn f s)
(h₂ : ∀ z ∈ s, f z ≠ 0) (h₂ : ∀ z ∈ s, f z ≠ 0)
(h₃ : ∀ z ∈ s, f z ∈ Complex.slitPlane) : (h₃ : ∀ z ∈ s, f z ∈ Complex.slitPlane) :
HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s := by HarmonicOn (Real.log ∘ Complex.normSq ∘ f) s := by
suffices hyp : Harmonic (⇑Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f) from suffices hyp : HarmonicOn (⇑Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f) s from
(harmonic_comp_CLM_is_harmonic hyp : Harmonic (Complex.reCLM ∘ Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f)) (harmonicOn_comp_CLM_is_harmonicOn hs hyp : HarmonicOn (Complex.reCLM ∘ Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f) s)
suffices hyp : Harmonic (Complex.log ∘ (((starRingEnd ) ∘ f) * f)) from by suffices hyp : HarmonicOn (Complex.log ∘ (((starRingEnd ) ∘ f) * f)) s from by
have : Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f = Complex.log ∘ (((starRingEnd ) ∘ f) * f) := by have : Complex.ofRealCLM ∘ Real.log ∘ Complex.normSq ∘ f = Complex.log ∘ (((starRingEnd ) ∘ f) * f) := by
funext z funext z
simp simp
@ -220,16 +220,16 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
-- Suffices to show Harmonic (Complex.log ∘ ⇑(starRingEnd ) ∘ f + Complex.log ∘ f) -- Suffices to show Harmonic (Complex.log ∘ ⇑(starRingEnd ) ∘ f + Complex.log ∘ f)
-- THIS IS WHERE WE USE h₃ -- THIS IS WHERE WE USE h₃
have : Complex.log ∘ (⇑(starRingEnd ) ∘ f * f) = Complex.log ∘ ⇑(starRingEnd ) ∘ f + Complex.log ∘ f := by have : ∀ z ∈ s, (Complex.log ∘ (⇑(starRingEnd ) ∘ f * f)) z = (Complex.log ∘ ⇑(starRingEnd ) ∘ f + Complex.log ∘ f) z := by
intro z hz
unfold Function.comp unfold Function.comp
funext z
simp simp
rw [Complex.log_mul_eq_add_log_iff] rw [Complex.log_mul_eq_add_log_iff]
have : Complex.arg ((starRingEnd ) (f z)) = - Complex.arg (f z) := by have : Complex.arg ((starRingEnd ) (f z)) = - Complex.arg (f z) := by
rw [Complex.arg_conj] rw [Complex.arg_conj]
have : ¬ Complex.arg (f z) = Real.pi := by have : ¬ Complex.arg (f z) = Real.pi := by
exact Complex.slitPlane_arg_ne_pi (h₃ z) exact Complex.slitPlane_arg_ne_pi (h₃ z hz)
simp simp
tauto tauto
rw [this] rw [this]
@ -237,8 +237,9 @@ theorem log_normSq_of_holomorphicOn_is_harmonicOn
constructor constructor
· exact Real.pi_pos · exact Real.pi_pos
· exact Real.pi_nonneg · exact Real.pi_nonneg
exact (AddEquivClass.map_ne_zero_iff starRingAut).mpr (h₂ z) exact (AddEquivClass.map_ne_zero_iff starRingAut).mpr (h₂ z hz)
exact h₂ z exact h₂ z hz
rw [this] rw [this]
apply harmonic_add_harmonic_is_harmonic apply harmonic_add_harmonic_is_harmonic