Update complexHarmonic.lean
This commit is contained in:
parent
9ac79470cd
commit
dc544308c8
|
@ -7,14 +7,49 @@ import Mathlib.Analysis.Calculus.FDeriv.Basic
|
||||||
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
|
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
|
||||||
import Nevanlinna.cauchyRiemann
|
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
|
noncomputable def Complex.laplace : (ℂ → ℂ) → (ℂ → ℂ) := by
|
||||||
intro f
|
intro f
|
||||||
|
let fx := Real.partialDeriv 1 f
|
||||||
let fx := fun w ↦ (fderiv ℝ f w) 1
|
let fxx := Real.partialDeriv 1 fx
|
||||||
let fxx := fun w ↦ (fderiv ℝ fx w) 1
|
let fy := Real.partialDeriv Complex.I f
|
||||||
let fy := fun w ↦ (fderiv ℝ f w) Complex.I
|
let fyy := Real.partialDeriv Complex.I fy
|
||||||
let fyy := fun w ↦ (fderiv ℝ fy w) Complex.I
|
exact fxx + fyy
|
||||||
exact fun z ↦ (fxx z) + (fyy z)
|
|
||||||
|
|
||||||
|
|
||||||
def Harmonic (f : ℂ → ℂ) : Prop :=
|
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
|
∀ 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
|
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₀ : ∀ y, HasFDerivAt f (f' y) y := by
|
||||||
have h : Differentiable ℝ f := by
|
have h : Differentiable ℝ f := by
|
||||||
exact (contDiff_succ_iff_fderiv.1 hf).left
|
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) :=
|
have f'_is_differentiable : Differentiable ℝ (fderiv ℝ f) :=
|
||||||
(contDiff_succ_iff_fderiv.1 f'_is_real_C1).left
|
(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
|
constructor
|
||||||
· -- f is two times real continuously differentiable
|
· -- f is two times real continuously differentiable
|
||||||
|
@ -76,38 +116,25 @@ theorem holomorphic_is_harmonic {f : ℂ → ℂ} (h : Differentiable ℂ f) :
|
||||||
intro z
|
intro z
|
||||||
unfold Complex.laplace
|
unfold Complex.laplace
|
||||||
|
|
||||||
simp
|
rw [CauchyRiemann₄ h]
|
||||||
|
|
||||||
conv =>
|
|
||||||
left
|
|
||||||
right
|
|
||||||
arg 1
|
|
||||||
arg 2
|
|
||||||
intro z
|
|
||||||
rw [CauchyRiemann₁ (h z)]
|
|
||||||
|
|
||||||
|
|
||||||
have t₀ : ∀ z, DifferentiableAt ℝ (fun w ↦ (fderiv ℝ f w) 1) z := by
|
|
||||||
intro z
|
have t₁a : (fderiv ℝ (fun w ↦ Complex.I * (fderiv ℝ f w) 1) z)
|
||||||
let A := f'_is_differentiable
|
= Complex.I • (fderiv ℝ f_1 z) := by
|
||||||
|
rw [fderiv_const_mul]
|
||||||
fun_prop
|
fun_prop
|
||||||
|
|
||||||
have t₁ : ∀ x, (fderiv ℝ (fun w ↦ Complex.I * (fderiv ℝ f w) 1) z) x
|
rw [t₁a]
|
||||||
= Complex.I * ((fderiv ℝ (fun w ↦ (fderiv ℝ f w) 1) z) x) := by
|
|
||||||
intro x
|
|
||||||
rw [fderiv_const_mul]
|
|
||||||
simp
|
|
||||||
exact t₀ z
|
|
||||||
rw [t₁]
|
|
||||||
|
|
||||||
have t₂ : (fderiv ℝ (fun w => (fderiv ℝ f w) 1) z) Complex.I
|
have t₂ : (fderiv ℝ f_1 z) Complex.I = (fderiv ℝ f_I z) 1 := by
|
||||||
= (fderiv ℝ (fun w => (fderiv ℝ f w) Complex.I) z) 1 := by
|
let B := l₂ f_is_real_C2 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 [← B]
|
||||||
|
let A := derivSymm f f_is_real_C2 z 1 Complex.I
|
||||||
rw [A]
|
rw [A]
|
||||||
let C := l₂ f_is_real_C2 z 1 Complex.I
|
let C := l₂ f_is_real_C2 z 1 Complex.I
|
||||||
rw [C]
|
rw [C]
|
||||||
|
simp
|
||||||
rw [t₂]
|
rw [t₂]
|
||||||
|
|
||||||
conv =>
|
conv =>
|
||||||
|
@ -117,9 +144,11 @@ theorem holomorphic_is_harmonic {f : ℂ → ℂ} (h : Differentiable ℂ f) :
|
||||||
arg 1
|
arg 1
|
||||||
arg 2
|
arg 2
|
||||||
intro z
|
intro z
|
||||||
|
simp [f_I]
|
||||||
rw [CauchyRiemann₁ (h z)]
|
rw [CauchyRiemann₁ (h z)]
|
||||||
|
|
||||||
rw [t₁]
|
rw [t₁a]
|
||||||
|
|
||||||
|
simp
|
||||||
rw [← mul_assoc]
|
rw [← mul_assoc]
|
||||||
simp
|
simp
|
||||||
|
|
Loading…
Reference in New Issue