Update complexHarmonic.lean

This commit is contained in:
Stefan Kebekus 2024-05-06 17:01:10 +02:00
parent b03954eee9
commit 9ac79470cd
1 changed files with 22 additions and 17 deletions

View File

@ -52,14 +52,25 @@ lemma l₂ {f : } (hf : ContDiff 2 f) (z a b : ) :
· simp
theorem holomorphic_is_harmonic (f : ) :
Differentiable f → Harmonic f := by
theorem holomorphic_is_harmonic {f : } (h : Differentiable f) :
Harmonic f := by
-- f is real C²
have f_is_real_C2 : ContDiff 2 f :=
ContDiff.restrict_scalars (Differentiable.contDiff h)
-- f' is real C¹
have f'_is_real_C1 : ContDiff 1 (fderiv f) :=
(contDiff_succ_iff_fderiv.1 f_is_real_C2).right
-- f' is real differentiable
have f'_is_differentiable : Differentiable (fderiv f) :=
(contDiff_succ_iff_fderiv.1 f'_is_real_C1).left
intro h
constructor
· -- f is two times real continuously differentiable
exact ContDiff.restrict_scalars (Differentiable.contDiff h)
exact f_is_real_C2
· -- Laplace of f is zero
intro z
@ -75,20 +86,14 @@ theorem holomorphic_is_harmonic (f : ) :
intro z
rw [CauchyRiemann₁ (h z)]
have t₂₀ : ContDiff 2 f := by exact ContDiff.restrict_scalars (Differentiable.contDiff h)
have t₀₀ : Differentiable (fun w => (fderiv f w)) := by
let A := (contDiff_succ_iff_fderiv.1 t₂₀).right
let B := (contDiff_succ_iff_fderiv.1 A).left
exact B
have t₀ : ∀ z, DifferentiableAt (fun w => (fderiv f w) 1) z := by
have t₀ : ∀ z, DifferentiableAt (fun w ↦ (fderiv f w) 1) z := by
intro z
let A := t₀₀
let A := f'_is_differentiable
fun_prop
have t₁ : ∀ x, (fderiv (fun w => Complex.I * (fderiv f w) 1) z) x
= Complex.I * ((fderiv (fun w => (fderiv f w) 1) z) x) := by
have t₁ : ∀ x, (fderiv (fun w ↦ Complex.I * (fderiv f w) 1) z) x
= Complex.I * ((fderiv (fun w (fderiv f w) 1) z) x) := by
intro x
rw [fderiv_const_mul]
simp
@ -97,11 +102,11 @@ theorem holomorphic_is_harmonic (f : ) :
have t₂ : (fderiv (fun w => (fderiv f w) 1) z) Complex.I
= (fderiv (fun w => (fderiv f w) Complex.I) z) 1 := by
let A := derivSymm f t₂₀ z 1 Complex.I
let B := l₂ t₂₀ z Complex.I 1
let A := derivSymm f f_is_real_C2 z 1 Complex.I
let B := l₂ f_is_real_C2 z Complex.I 1
rw [← B]
rw [A]
let C := l₂ t₂₀ z 1 Complex.I
let C := l₂ f_is_real_C2 z 1 Complex.I
rw [C]
rw [t₂]