nevanlinna/Nevanlinna/holomorphic.lean

136 lines
3.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 Mathlib.Analysis.Calculus.FDeriv.Symmetric
import Mathlib.Analysis.RCLike.Basic
import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv
import Mathlib.Data.Complex.Module
import Mathlib.Data.Complex.Order
import Mathlib.Data.Complex.Exponential
import Mathlib.Data.Fin.Tuple.Basic
import Mathlib.Topology.Algebra.InfiniteSum.Module
import Mathlib.Topology.Defs.Filter
import Mathlib.Topology.Instances.RealVectorSpace
import Nevanlinna.cauchyRiemann
import Nevanlinna.laplace
import Nevanlinna.complexHarmonic
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace E]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace F] [CompleteSpace F]
def HolomorphicAt (f : E → F) (x : E) : Prop :=
∃ s ∈ nhds x, ∀ z ∈ s, DifferentiableAt f z
theorem HolomorphicAt_iff
{f : E → F}
{x : E} :
HolomorphicAt f x ↔ ∃ s : Set E, IsOpen s ∧ x ∈ s ∧ (∀ z ∈ s, DifferentiableAt f z) := by
constructor
· intro hf
obtain ⟨t, h₁t, h₂t⟩ := hf
obtain ⟨s, h₁s, h₂s, h₃s⟩ := mem_nhds_iff.1 h₁t
use s
constructor
· assumption
· constructor
· assumption
· intro z hz
exact h₂t z (h₁s hz)
· intro hyp
obtain ⟨s, h₁s, h₂s, hf⟩ := hyp
use s
constructor
· apply (IsOpen.mem_nhds_iff h₁s).2 h₂s
· assumption
theorem HolomorphicAt_differentiableAt
{f : E → F}
{x : E} :
HolomorphicAt f x → DifferentiableAt f x := by
intro hf
obtain ⟨s, _, h₂s, h₃s⟩ := HolomorphicAt_iff.1 hf
exact h₃s x h₂s
theorem HolomorphicAt_isOpen
(f : E → F) :
IsOpen { x : E | HolomorphicAt f x } := by
rw [← subset_interior_iff_isOpen]
intro x hx
simp at hx
obtain ⟨s, h₁s, h₂s, h₃s⟩ := HolomorphicAt_iff.1 hx
use s
constructor
· simp
constructor
· exact h₁s
· intro x hx
simp
use s
constructor
· exact IsOpen.mem_nhds h₁s hx
· exact h₃s
· exact h₂s
theorem HolomorphicAt_contDiffAt
{f : → F}
{z : }
(hf : HolomorphicAt f z) :
ContDiffAt 2 f z := by
let t := {x | HolomorphicAt f x}
have ht : IsOpen t := HolomorphicAt_isOpen f
have hz : z ∈ t := by exact hf
-- ContDiffAt 2 f z
apply ContDiffOn.contDiffAt _ ((IsOpen.mem_nhds_iff ht).2 hz)
suffices h : ContDiffOn 2 f t from by
apply ContDiffOn.restrict_scalars h
apply DifferentiableOn.contDiffOn _ ht
intro w hw
apply DifferentiableAt.differentiableWithinAt
exact HolomorphicAt_differentiableAt hw
theorem CauchyRiemann'₅
{f : → F}
{z : }
(h : DifferentiableAt f z) :
partialDeriv Complex.I f z = Complex.I • partialDeriv 1 f z := by
unfold partialDeriv
conv =>
left
rw [DifferentiableAt.fderiv_restrictScalars h]
simp
rw [← mul_one Complex.I]
rw [← smul_eq_mul]
rw [ContinuousLinearMap.map_smul_of_tower (fderiv f z) Complex.I 1]
conv =>
right
right
rw [DifferentiableAt.fderiv_restrictScalars h]
theorem CauchyRiemann'₆
{f : → F}
{z : }
(h : HolomorphicAt f z) :
partialDeriv Complex.I f =ᶠ[nhds z] Complex.I • partialDeriv 1 f := by
obtain ⟨s, h₁s, hz, h₂f⟩ := HolomorphicAt_iff.1 h
apply Filter.eventuallyEq_iff_exists_mem.2
use s
constructor
· exact IsOpen.mem_nhds h₁s hz
· intro w hw
apply CauchyRiemann'₅
exact h₂f w hw