Update complexHarmonic.lean

This commit is contained in:
Stefan Kebekus 2024-05-03 12:23:09 +02:00
parent ea3693ff24
commit bfdc7f6d2a
1 changed files with 53 additions and 11 deletions

View File

@ -9,16 +9,55 @@ import Nevanlinna.cauchyRiemann
noncomputable def Complex.laplace : () → () := by
intro f
let fx := fun w ↦ fderiv f w 1
let fxx := fun z ↦ fderiv fx z 1
let fy := fun w ↦ fderiv f w Complex.I
let fyy := fun z ↦ fderiv fy z Complex.I
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)
def Harmonic (f : ) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
#check second_derivative_symmetric
lemma zwoDiff (f : × ) (h : ContDiff 2 f) : ∀ z a b : , 0 = 1 := by
intro z a b
let fx := fun w ↦ (fderiv f w) 1
let fxx := fun w ↦ (fderiv fx w) 1
let f2 := (fderiv (fun w => fderiv f w) z) 1 1
have : iteratedFDeriv (1 + 1) f = 0 := by
rw [iteratedFDeriv_succ_eq_comp_left]
simp
sorry
have : f2 = fxx z := by
dsimp [f2, fxx, fx]
sorry
sorry
lemma derivSymm (f : ) (h : Differentiable 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)
have h₀ : ∀ y, HasFDerivAt f (f' y) y := by
exact fun y => DifferentiableAt.hasFDerivAt (h y)
let f'' := (fderiv f' z)
have h₁ : HasFDerivAt f' f'' z := by
apply DifferentiableAt.hasFDerivAt
sorry
let A := second_derivative_symmetric h₀ h₁ a b
dsimp [f'', f'] at A
apply A
theorem re_comp_holomorphic_is_harmonic (f : ) :
Differentiable f → Harmonic f := by
@ -32,6 +71,7 @@ theorem re_comp_holomorphic_is_harmonic (f : ) :
· -- Laplace of f is zero
intro z
unfold Complex.laplace
simp
conv =>
@ -41,10 +81,10 @@ theorem re_comp_holomorphic_is_harmonic (f : ) :
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
have t₀ : ∀ z, DifferentiableAt (fun w => (fderiv f w) 1) z := by
intro z
sorry
have t₁ : ∀ x, (fderiv (fun w => Complex.I * (fderiv f w) 1) z) x
@ -54,9 +94,12 @@ theorem re_comp_holomorphic_is_harmonic (f : ) :
simp
exact t₀ z
rw [t₁]
have t₂ : (fderiv (fun w => (fderiv f w) 1) z) Complex.I
have t₂₀ : Differentiable f := by sorry
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
sorry
rw [t₂]
@ -68,9 +111,8 @@ theorem re_comp_holomorphic_is_harmonic (f : ) :
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
rw [t₁]
rw [← mul_assoc]
simp