nevanlinna/Nevanlinna/complexHarmonic.lean

53 lines
1.6 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.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 Nevanlinna.cauchyRiemann
noncomputable def Complex.laplace : () → () := by
intro f
let f₁ := fun x ↦ lineDeriv f x 1
let f₁₁ := fun x ↦ lineDeriv f₁ x 1
let f₂ := fun x ↦ lineDeriv f x Complex.I
let f₂₂ := fun x ↦ lineDeriv f₂ x Complex.I
exact f₁₁ + f₂₂
def Harmonic (f : ) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
#check contDiff_iff_ftaylorSeries.2
lemma c2_if_holomorphic (f : ) : Differentiable f → ContDiff 2 f := by
intro fHyp
exact Differentiable.contDiff fHyp
lemma c2R_if_holomorphic (f : ) : Differentiable f → ContDiff 2 f := by
intro fHyp
let ZZ := c2_if_holomorphic f fHyp
apply ContDiff.restrict_scalars ZZ
theorem re_comp_holomorphic_is_harmonic (f : ) :
Differentiable f → Harmonic (Complex.reCLM ∘ f) := by
intro h
constructor
· -- Complex.reCLM ∘ f is two times real continuously differentiable
apply ContDiff.comp
· -- Complex.reCLM is two times real continuously differentiable
exact ContinuousLinearMap.contDiff Complex.reCLM
· -- f is two times real continuously differentiable
exact c2R_if_holomorphic f h
· -- Laplace of f is zero
intro z
unfold Complex.laplace
simp
let ZZ := (CauchyRiemann₃ (h z)).left
sorry