nevanlinna/Nevanlinna/complexHarmonic.lean

126 lines
3.4 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Data.Fin.Tuple.Basic
import Mathlib.Analysis.Complex.Basic
import Mathlib.Analysis.Complex.TaylorSeries
import Mathlib.Analysis.Calculus.LineDeriv.Basic
import Mathlib.Analysis.Calculus.ContDiff.Defs
import Mathlib.Analysis.Calculus.FDeriv.Basic
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
import Nevanlinna.cauchyRiemann
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)
def Harmonic (f : ) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
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)
have h₀ : ∀ y, HasFDerivAt f (f' y) y := by
have h : Differentiable f := by
exact (contDiff_succ_iff_fderiv.1 hf).left
exact fun y => DifferentiableAt.hasFDerivAt (h y)
let f'' := (fderiv f' z)
have h₁ : HasFDerivAt f' f'' z := by
apply DifferentiableAt.hasFDerivAt
let A := (contDiff_succ_iff_fderiv.1 hf).right
let B := (contDiff_succ_iff_fderiv.1 A).left
simp at B
exact B z
let A := second_derivative_symmetric h₀ h₁ a b
dsimp [f'', f'] at A
apply A
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
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
constructor
· -- f is two times real continuously differentiable
exact f_is_real_C2
· -- Laplace of f is zero
intro z
unfold Complex.laplace
simp
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
let A := f'_is_differentiable
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₁]
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
let B := l₂ f_is_real_C2 z Complex.I 1
rw [← B]
rw [A]
let C := l₂ f_is_real_C2 z 1 Complex.I
rw [C]
rw [t₂]
conv =>
left
right
arg 2
arg 1
arg 2
intro z
rw [CauchyRiemann₁ (h z)]
rw [t₁]
rw [← mul_assoc]
simp