Update complexHarmonic.lean

This commit is contained in:
Stefan Kebekus 2024-05-07 07:08:23 +02:00
parent 9ac79470cd
commit dc544308c8
1 changed files with 61 additions and 32 deletions

View File

@ -7,14 +7,49 @@ import Mathlib.Analysis.Calculus.FDeriv.Basic
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
import Nevanlinna.cauchyRiemann
noncomputable def Real.partialDeriv : → () → () := by
intro v
intro f
exact fun w ↦ (fderiv f w) v
theorem CauchyRiemann₄ {f : } : (Differentiable f)
→ Real.partialDeriv Complex.I f = Complex.I • Real.partialDeriv 1 f := by
intro h
unfold Real.partialDeriv
conv =>
left
intro w
rw [DifferentiableAt.fderiv_restrictScalars (h w)]
simp
rw [← mul_one Complex.I]
rw [← smul_eq_mul]
rw [ContinuousLinearMap.map_smul_of_tower (fderiv f w) Complex.I 1]
conv =>
right
right
intro w
rw [DifferentiableAt.fderiv_restrictScalars (h w)]
theorem partialDeriv_smul {f : } {a v : } : Real.partialDeriv v (a • f) = a • Real.partialDeriv v f := by
unfold Real.partialDeriv
have : a • f = fun y ↦ a • f y := by rfl
conv =>
left
intro w
rw [this]
rw [fderiv_const_smul]
sorry
noncomputable def Complex.laplace : () → () := by
intro f
let fx := fun w ↦ (fderiv f w) 1
let fxx := fun w ↦ (fderiv fx w) 1
let fy := fun w ↦ (fderiv f w) Complex.I
let fyy := fun w ↦ (fderiv fy w) Complex.I
exact fun z ↦ (fxx z) + (fyy z)
let fx := Real.partialDeriv 1 f
let fxx := Real.partialDeriv 1 fx
let fy := Real.partialDeriv Complex.I f
let fyy := Real.partialDeriv Complex.I fy
exact fxx + fyy
def Harmonic (f : ) : Prop :=
@ -25,7 +60,7 @@ lemma derivSymm (f : ) (hf : ContDiff 2 f) :
∀ z a b : , (fderiv (fun w => fderiv f w) z) a b = (fderiv (fun w => fderiv f w) z) b a := by
intro z a b
let f' := fun w => (fderiv f w)
let f' := fderiv f
have h₀ : ∀ y, HasFDerivAt f (f' y) y := by
have h : Differentiable f := by
exact (contDiff_succ_iff_fderiv.1 hf).left
@ -67,6 +102,11 @@ theorem holomorphic_is_harmonic {f : } (h : Differentiable f) :
have f'_is_differentiable : Differentiable (fderiv f) :=
(contDiff_succ_iff_fderiv.1 f'_is_real_C1).left
-- Partial derivative in direction 1
let f_1 := fun w ↦ (fderiv f w) 1
-- Partial derivative in direction I
let f_I := fun w ↦ (fderiv f w) Complex.I
constructor
· -- f is two times real continuously differentiable
@ -76,38 +116,25 @@ theorem holomorphic_is_harmonic {f : } (h : Differentiable f) :
intro z
unfold Complex.laplace
simp
conv =>
left
right
arg 1
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
rw [CauchyRiemann₄ h]
have t₀ : ∀ z, DifferentiableAt (fun w ↦ (fderiv f w) 1) z := by
intro z
let A := f'_is_differentiable
have t₁a : (fderiv (fun w ↦ Complex.I * (fderiv f w) 1) z)
= Complex.I • (fderiv f_1 z) := by
rw [fderiv_const_mul]
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
intro x
rw [fderiv_const_mul]
simp
exact t₀ z
rw [t₁]
rw [t₁a]
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 f_is_real_C2 z 1 Complex.I
have t₂ : (fderiv f_1 z) Complex.I = (fderiv f_I z) 1 := by
let B := l₂ f_is_real_C2 z Complex.I 1
rw [← B]
let A := derivSymm f f_is_real_C2 z 1 Complex.I
rw [A]
let C := l₂ f_is_real_C2 z 1 Complex.I
rw [C]
simp
rw [t₂]
conv =>
@ -117,9 +144,11 @@ theorem holomorphic_is_harmonic {f : } (h : Differentiable f) :
arg 1
arg 2
intro z
simp [f_I]
rw [CauchyRiemann₁ (h z)]
rw [t₁]
rw [t₁a]
simp
rw [← mul_assoc]
simp