nevanlinna/Nevanlinna/holomorphic.lean

102 lines
2.6 KiB
Plaintext
Raw Normal View History

2024-06-05 10:19:09 +02:00
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
2024-06-05 12:03:05 +02:00
theorem HolomorphicAt_isOpen
{f : E → F}
{x : E} :
HolomorphicAt f x ↔ ∃ s : Set E, IsOpen s ∧ x ∈ s ∧ (∀ z ∈ s, HolomorphicAt 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
apply HolomorphicAt_iff.2
use s
constructor
· assumption
· constructor
· assumption
· exact fun w hw ↦ h₂t w (h₁s hw)
· intro hyp
obtain ⟨s, h₁s, h₂s, hf⟩ := hyp
use s
constructor
· apply (IsOpen.mem_nhds_iff h₁s).2 h₂s
· intro z hz
obtain ⟨t, ht⟩ := (hf z hz)
exact ht.2 z (mem_of_mem_nhds ht.1)
theorem CauchyRiemann'₅
2024-06-05 10:19:09 +02:00
{f : → F}
2024-06-05 12:03:05 +02:00
{z : }
(h : DifferentiableAt f z) :
partialDeriv Complex.I f z = Complex.I • partialDeriv 1 f z := by
2024-06-05 10:19:09 +02:00
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]
2024-06-05 12:03:05 +02:00
theorem CauchyRiemann'₆
{f : → F}
{z : }
(h : HolomorphicAt f z) :
partialDeriv Complex.I f =ᶠ[nhds z] Complex.I • partialDeriv 1 f := by
2024-06-05 10:19:09 +02:00
2024-06-05 12:03:05 +02:00
obtain ⟨s, h₁s, hz, h₂f⟩ := HolomorphicAt_iff.1 h
2024-06-05 10:19:09 +02:00
2024-06-05 12:03:05 +02:00
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