Update complexHarmonic.lean

This commit is contained in:
Stefan Kebekus 2024-05-14 17:24:34 +02:00
parent 06f8de2f56
commit a60e07e602
1 changed files with 6 additions and 6 deletions

View File

@ -142,18 +142,18 @@ theorem logabs_of_holomorphic_is_harmonic
constructor constructor
· -- logabs f is real C² · -- logabs f is real C²
have : (fun z ↦ Real.log ‖f z‖) = (Real.log ∘ Complex.normSq ∘ f) / 2 := by have : (fun z ↦ Real.log ‖f z‖) = (2 : )⁻¹ • (Real.log ∘ Complex.normSq ∘ f) := by
funext z funext z
simp simp
unfold Complex.abs unfold Complex.abs
simp simp
rw [Real.log_sqrt] rw [Real.log_sqrt]
exact Complex.normSq_nonneg (f z) rw [div_eq_inv_mul (Real.log (Complex.normSq (f z))) 2]
exact Complex.normSq_nonneg (f z)
rw [this] rw [this]
have : Real.log ∘ ⇑Complex.normSq ∘ f / 2 = (fun z ↦ (1 / (2 : )) • ((Real.log ∘ ⇑Complex.normSq ∘ f) z)) := by
funext z have : (2 : )⁻¹ • (Real.log ∘ Complex.normSq ∘ f) = (fun z ↦ (2 : )⁻¹ • ((Real.log ∘ ⇑Complex.normSq ∘ f) z)) := by
simp simp
exact div_eq_inv_mul (Real.log (Complex.normSq (f z))) 2
rw [this] rw [this]
apply contDiff_iff_contDiffAt.2 apply contDiff_iff_contDiffAt.2
@ -169,5 +169,5 @@ theorem logabs_of_holomorphic_is_harmonic
exact ContDiff.contDiffAt f_is_real_C2 exact ContDiff.contDiffAt f_is_real_C2
· -- Laplace vanishes · -- Laplace vanishes
sorry sorry