2024-05-03 15:54:51 +02:00
|
|
|
|
import Mathlib.Data.Fin.Tuple.Basic
|
2024-04-30 08:20:57 +02:00
|
|
|
|
import Mathlib.Analysis.Complex.Basic
|
2024-05-02 09:48:26 +02:00
|
|
|
|
import Mathlib.Analysis.Complex.TaylorSeries
|
2024-04-30 08:20:57 +02:00
|
|
|
|
import Mathlib.Analysis.Calculus.LineDeriv.Basic
|
|
|
|
|
import Mathlib.Analysis.Calculus.ContDiff.Defs
|
|
|
|
|
import Mathlib.Analysis.Calculus.FDeriv.Basic
|
2024-05-02 21:09:38 +02:00
|
|
|
|
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
|
2024-04-30 08:20:57 +02:00
|
|
|
|
import Nevanlinna.cauchyRiemann
|
2024-05-07 09:49:56 +02:00
|
|
|
|
import Nevanlinna.partialDeriv
|
2024-05-07 07:08:23 +02:00
|
|
|
|
|
|
|
|
|
theorem CauchyRiemann₄ {f : ℂ → ℂ} : (Differentiable ℂ f)
|
|
|
|
|
→ Real.partialDeriv Complex.I f = Complex.I • Real.partialDeriv 1 f := by
|
|
|
|
|
intro h
|
|
|
|
|
unfold Real.partialDeriv
|
2024-05-07 09:49:56 +02:00
|
|
|
|
|
2024-05-07 07:08:23 +02:00
|
|
|
|
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
|
2024-05-07 09:49:56 +02:00
|
|
|
|
intro w
|
2024-05-07 07:08:23 +02:00
|
|
|
|
rw [DifferentiableAt.fderiv_restrictScalars ℝ (h w)]
|
|
|
|
|
|
2024-04-30 08:20:57 +02:00
|
|
|
|
|
2024-05-07 07:08:23 +02:00
|
|
|
|
noncomputable def Complex.laplace : (ℂ → ℂ) → (ℂ → ℂ) := by
|
|
|
|
|
intro f
|
|
|
|
|
let fx := Real.partialDeriv 1 f
|
|
|
|
|
let fxx := Real.partialDeriv 1 fx
|
2024-05-07 09:49:56 +02:00
|
|
|
|
let fy := Real.partialDeriv Complex.I f
|
2024-05-07 07:08:23 +02:00
|
|
|
|
let fyy := Real.partialDeriv Complex.I fy
|
|
|
|
|
exact fxx + fyy
|
2024-04-30 08:20:57 +02:00
|
|
|
|
|
2024-05-02 21:09:38 +02:00
|
|
|
|
|
|
|
|
|
def Harmonic (f : ℂ → ℂ) : Prop :=
|
2024-04-30 08:20:57 +02:00
|
|
|
|
(ContDiff ℝ 2 f) ∧ (∀ z, Complex.laplace f z = 0)
|
|
|
|
|
|
2024-05-03 12:23:09 +02:00
|
|
|
|
|
2024-05-06 10:09:49 +02:00
|
|
|
|
lemma derivSymm (f : ℂ → ℂ) (hf : ContDiff ℝ 2 f) :
|
2024-05-03 12:23:09 +02:00
|
|
|
|
∀ 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
|
|
|
|
|
|
2024-05-07 07:08:23 +02:00
|
|
|
|
let f' := fderiv ℝ f
|
2024-05-03 12:23:09 +02:00
|
|
|
|
have h₀ : ∀ y, HasFDerivAt f (f' y) y := by
|
2024-05-06 10:09:49 +02:00
|
|
|
|
have h : Differentiable ℝ f := by
|
|
|
|
|
exact (contDiff_succ_iff_fderiv.1 hf).left
|
2024-05-03 12:23:09 +02:00
|
|
|
|
exact fun y => DifferentiableAt.hasFDerivAt (h y)
|
|
|
|
|
|
|
|
|
|
let f'' := (fderiv ℝ f' z)
|
|
|
|
|
have h₁ : HasFDerivAt f' f'' z := by
|
|
|
|
|
apply DifferentiableAt.hasFDerivAt
|
2024-05-06 10:09:49 +02:00
|
|
|
|
let A := (contDiff_succ_iff_fderiv.1 hf).right
|
|
|
|
|
let B := (contDiff_succ_iff_fderiv.1 A).left
|
|
|
|
|
simp at B
|
|
|
|
|
exact B z
|
2024-05-03 12:23:09 +02:00
|
|
|
|
|
|
|
|
|
let A := second_derivative_symmetric h₀ h₁ a b
|
|
|
|
|
dsimp [f'', f'] at A
|
|
|
|
|
apply A
|
2024-05-02 09:48:26 +02:00
|
|
|
|
|
2024-05-06 09:01:43 +02:00
|
|
|
|
|
2024-05-06 10:09:49 +02:00
|
|
|
|
lemma l₂ {f : ℂ → ℂ} (hf : ContDiff ℝ 2 f) (z a b : ℂ) :
|
|
|
|
|
fderiv ℝ (fderiv ℝ f) z b a = fderiv ℝ (fun w ↦ fderiv ℝ f w a) z b := by
|
|
|
|
|
rw [fderiv_clm_apply]
|
|
|
|
|
· simp
|
|
|
|
|
· exact (contDiff_succ_iff_fderiv.1 hf).2.differentiable le_rfl z
|
|
|
|
|
· simp
|
|
|
|
|
|
|
|
|
|
|
2024-05-06 17:01:10 +02:00
|
|
|
|
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)
|
|
|
|
|
|
2024-05-07 09:49:56 +02:00
|
|
|
|
have fI_is_real_differentiable : Differentiable ℝ (Real.partialDeriv 1 f) := by
|
2024-05-07 10:16:23 +02:00
|
|
|
|
exact (partialDeriv_contDiff f_is_real_C2 1).differentiable (Submonoid.oneLE.proof_2 ℕ∞)
|
2024-04-30 08:20:57 +02:00
|
|
|
|
|
|
|
|
|
constructor
|
2024-05-06 09:01:43 +02:00
|
|
|
|
· -- f is two times real continuously differentiable
|
2024-05-06 17:01:10 +02:00
|
|
|
|
exact f_is_real_C2
|
2024-04-30 08:20:57 +02:00
|
|
|
|
|
|
|
|
|
· -- Laplace of f is zero
|
|
|
|
|
unfold Complex.laplace
|
2024-05-07 07:08:23 +02:00
|
|
|
|
rw [CauchyRiemann₄ h]
|
2024-05-07 09:49:56 +02:00
|
|
|
|
rw [partialDeriv_smul fI_is_real_differentiable]
|
2024-05-07 10:16:23 +02:00
|
|
|
|
rw [partialDeriv_comm f_is_real_C2 Complex.I 1]
|
|
|
|
|
rw [CauchyRiemann₄ h]
|
|
|
|
|
rw [partialDeriv_smul fI_is_real_differentiable]
|
|
|
|
|
rw [← smul_assoc]
|
2024-05-02 21:09:38 +02:00
|
|
|
|
simp
|