nevanlinna/Nevanlinna/holomorphic.lean

64 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.Calculus.ContDiff.Defs
import Nevanlinna.partialDeriv
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace E]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 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 CauchyRiemann₅
{f : → F}
{z : } :
(DifferentiableAt f z) → partialDeriv Complex.I f z = Complex.I • partialDeriv 1 f z := by
intro h
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 : }
{z : } :
(HolomorphicAt f z) → partialDeriv Complex.I f =ᶠ[nhds z] Complex.I • partialDeriv 1 f := by
sorry