2024-06-06 14:45:29 +02:00
|
|
|
|
import Mathlib.Analysis.Complex.TaylorSeries
|
|
|
|
|
import Nevanlinna.cauchyRiemann
|
2024-06-05 10:19:09 +02:00
|
|
|
|
|
|
|
|
|
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E]
|
2024-08-23 09:26:58 +02:00
|
|
|
|
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F]
|
|
|
|
|
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ℂ G]
|
2024-06-05 10:19:09 +02:00
|
|
|
|
|
|
|
|
|
def HolomorphicAt (f : E → F) (x : E) : Prop :=
|
|
|
|
|
∃ s ∈ nhds x, ∀ z ∈ s, DifferentiableAt ℂ f z
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
theorem HolomorphicAt_iff
|
|
|
|
|
{f : E → F}
|
|
|
|
|
{x : E} :
|
2024-06-10 12:39:40 +02:00
|
|
|
|
HolomorphicAt f x ↔ ∃ s :
|
|
|
|
|
Set E, IsOpen s ∧ x ∈ s ∧ (∀ z ∈ s, DifferentiableAt ℂ f z) := by
|
2024-06-05 10:19:09 +02:00
|
|
|
|
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-08-22 14:21:14 +02:00
|
|
|
|
theorem HolomorphicAt_analyticAt
|
2024-08-23 09:26:58 +02:00
|
|
|
|
[CompleteSpace F]
|
2024-08-22 14:21:14 +02:00
|
|
|
|
{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-06 14:45:29 +02:00
|
|
|
|
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) :
|
2024-06-06 09:52:42 +02:00
|
|
|
|
IsOpen { x : E | HolomorphicAt f x } := by
|
2024-06-10 10:58:57 +02:00
|
|
|
|
|
2024-06-06 09:52:42 +02:00
|
|
|
|
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
|
|
|
|
|
|
2024-06-06 14:45:29 +02:00
|
|
|
|
|
2024-06-10 10:58:57 +02:00
|
|
|
|
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
|
|
|
|
|
|
|
|
|
|
|
2024-06-10 12:39:40 +02:00
|
|
|
|
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-06-06 14:45:29 +02:00
|
|
|
|
theorem HolomorphicAt_contDiffAt
|
2024-08-23 09:26:58 +02:00
|
|
|
|
[CompleteSpace F]
|
2024-06-06 14:45:29 +02:00
|
|
|
|
{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
|
|
|
|
|
|
2024-06-05 12:03:05 +02:00
|
|
|
|
|
|
|
|
|
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]
|
|
|
|
|
conv =>
|
|
|
|
|
right
|
|
|
|
|
right
|
|
|
|
|
rw [DifferentiableAt.fderiv_restrictScalars ℝ h]
|
2024-07-25 15:52:16 +02:00
|
|
|
|
simp
|
2024-06-05 10:19:09 +02:00
|
|
|
|
|
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
|