nevanlinna/Nevanlinna/complexHarmonic.lean

126 lines
3.4 KiB
Plaintext
Raw Normal View History

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-02 21:09:38 +02:00
noncomputable def Complex.laplace : () → () := by
2024-04-30 08:20:57 +02:00
intro f
2024-05-03 12:23:09 +02:00
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
2024-05-02 21:09:38 +02:00
exact fun z ↦ (fxx z) + (fyy z)
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
let f' := fun w => (fderiv f w)
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)
-- 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
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
intro z
unfold Complex.laplace
2024-05-03 12:23:09 +02:00
2024-04-30 08:20:57 +02:00
simp
2024-05-02 21:09:38 +02:00
conv =>
left
right
arg 1
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
2024-05-03 12:23:09 +02:00
2024-05-06 10:09:49 +02:00
2024-05-06 17:01:10 +02:00
have t₀ : ∀ z, DifferentiableAt (fun w ↦ (fderiv f w) 1) z := by
2024-05-02 21:09:38 +02:00
intro z
2024-05-06 17:01:10 +02:00
let A := f'_is_differentiable
2024-05-06 10:09:49 +02:00
fun_prop
2024-05-02 21:09:38 +02:00
2024-05-06 17:01:10 +02:00
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
2024-05-02 21:09:38 +02:00
intro x
rw [fderiv_const_mul]
simp
exact t₀ z
rw [t₁]
2024-05-03 12:23:09 +02:00
have t₂ : (fderiv (fun w => (fderiv f w) 1) z) Complex.I
2024-05-02 21:09:38 +02:00
= (fderiv (fun w => (fderiv f w) Complex.I) z) 1 := by
2024-05-06 17:01:10 +02:00
let A := derivSymm f f_is_real_C2 z 1 Complex.I
let B := l₂ f_is_real_C2 z Complex.I 1
2024-05-06 10:09:49 +02:00
rw [← B]
rw [A]
2024-05-06 17:01:10 +02:00
let C := l₂ f_is_real_C2 z 1 Complex.I
2024-05-06 10:09:49 +02:00
rw [C]
2024-05-02 21:09:38 +02:00
rw [t₂]
conv =>
left
right
arg 2
arg 1
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
2024-05-03 12:23:09 +02:00
2024-05-02 21:09:38 +02:00
rw [t₁]
2024-05-03 12:23:09 +02:00
2024-05-02 21:09:38 +02:00
rw [← mul_assoc]
simp