nevanlinna/Nevanlinna/holomorphicAt.lean

178 lines
4.1 KiB
Plaintext
Raw Normal View History

2024-08-22 14:21:14 +02:00
import Mathlib.Analysis.Complex.CauchyIntegral
2024-06-11 13:28:02 +02:00
import Nevanlinna.cauchyRiemann
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace E]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace F]
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace G]
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-07-31 15:59:43 +02:00
theorem Differentiable.holomorphicAt
{f : E → F}
(hf : Differentiable f)
{x : E} :
HolomorphicAt f x := by
apply HolomorphicAt_iff.2
use Set.univ
constructor
· exact isOpen_univ
· constructor
· exact trivial
· intro z _
exact hf z
2024-06-11 13:28:02 +02:00
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_comp
{g : E → F}
{f : F → G}
{z : E}
(hf : HolomorphicAt f (g z))
(hg : HolomorphicAt g z) :
HolomorphicAt (f ∘ g) z := by
obtain ⟨UE, h₁UE, h₂UE⟩ := hg
obtain ⟨UF, h₁UF, h₂UF⟩ := hf
use UE ∩ g⁻¹' UF
constructor
· simp
constructor
· assumption
· apply ContinuousAt.preimage_mem_nhds
apply (h₂UE z (mem_of_mem_nhds h₁UE)).continuousAt
assumption
· intro x hx
apply DifferentiableAt.comp
apply h₂UF
exact hx.2
apply h₂UE
exact hx.1
theorem HolomorphicAt_neg
{f : E → F}
{z : E}
(hf : HolomorphicAt f z) :
HolomorphicAt (-f) z := by
obtain ⟨UF, h₁UF, h₂UF⟩ := hf
use UF
constructor
· assumption
· intro z hz
apply differentiableAt_neg_iff.mp
simp
exact h₂UF z hz
2024-08-22 14:21:14 +02:00
theorem HolomorphicAt.analyticAt
[CompleteSpace F]
{f : → F}
{x : } :
HolomorphicAt f x → AnalyticAt f x := by
intro hf
obtain ⟨s, h₁s, h₂s, h₃s⟩ := HolomorphicAt_iff.1 hf
apply DifferentiableOn.analyticAt (s := s)
intro z hz
apply DifferentiableAt.differentiableWithinAt
apply h₃s
exact hz
exact IsOpen.mem_nhds h₁s h₂s
2024-06-11 13:28:02 +02:00
theorem HolomorphicAt.contDiffAt
[CompleteSpace F]
{f : → F}
{z : }
{n : }
(hf : HolomorphicAt f z) :
ContDiffAt n f z := by
let t := {x | HolomorphicAt f x}
have ht : IsOpen t := HolomorphicAt_isOpen f
have hz : z ∈ t := by exact hf
-- ContDiffAt _ f z
apply ContDiffOn.contDiffAt _ ((IsOpen.mem_nhds_iff ht).2 hz)
suffices h : ContDiffOn n f t from by
apply ContDiffOn.restrict_scalars h
apply DifferentiableOn.contDiffOn _ ht
intro w hw
apply DifferentiableAt.differentiableWithinAt
-- DifferentiableAt f w
let hfw : HolomorphicAt f w := hw
obtain ⟨s, _, h₂s, h₃s⟩ := HolomorphicAt_iff.1 hfw
exact h₃s w h₂s
theorem HolomorphicAt.differentiableAt
{f : → F}
{z : }
(hf : HolomorphicAt f z) :
DifferentiableAt f z := by
obtain ⟨s, _, h₂s, h₃s⟩ := HolomorphicAt_iff.1 hf
exact h₃s z h₂s
theorem HolomorphicAt.CauchyRiemannAt
{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
let h := h₂f w hw
unfold partialDeriv
2024-06-11 17:18:24 +02:00
apply CauchyRiemann₅ h