123 lines
3.0 KiB
Plaintext
123 lines
3.0 KiB
Plaintext
import Mathlib.Analysis.Complex.Basic
|
||
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 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_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'₅
|
||
{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
|