Working
This commit is contained in:
parent
a6aba0fc68
commit
2f4672e144
|
@ -30,19 +30,15 @@ theorem holomorphicAt_is_harmonicAt
|
|||
(hf : HolomorphicAt f z) :
|
||||
HarmonicAt f z := by
|
||||
|
||||
obtain ⟨s, hs, hz, h'f⟩ := HolomorphicAt_iff.1 hf
|
||||
let t := {x | HolomorphicAt f x}
|
||||
have ht : IsOpen t := HolomorphicAt_isOpen f
|
||||
have hz : z ∈ t := by exact hf
|
||||
|
||||
constructor
|
||||
· -- ContDiffAt ℝ 2 f z
|
||||
apply ContDiffOn.contDiffAt _ ((IsOpen.mem_nhds_iff hs).2 hz)
|
||||
suffices h : ContDiffOn ℂ 2 f s from by
|
||||
apply ContDiffOn.restrict_scalars ℝ h
|
||||
apply DifferentiableOn.contDiffOn _ hs
|
||||
intro w hw
|
||||
apply DifferentiableAt.differentiableWithinAt
|
||||
exact h'f w hw
|
||||
exact HolomorphicAt_contDiffAt hf
|
||||
|
||||
· -- Δ f =ᶠ[nhds z] 0
|
||||
obtain ⟨t, ht, hz, h'f⟩ := HolomorphicAt_isOpen.1 hf
|
||||
apply Filter.eventuallyEq_iff_exists_mem.2
|
||||
use t
|
||||
constructor
|
||||
|
@ -50,15 +46,12 @@ theorem holomorphicAt_is_harmonicAt
|
|||
· intro w hw
|
||||
unfold Complex.laplace
|
||||
simp
|
||||
|
||||
rw [partialDeriv_eventuallyEq ℝ (CauchyRiemann'₆ (h'f w hw)) Complex.I]
|
||||
rw [partialDeriv_eventuallyEq ℝ (CauchyRiemann'₆ hw) Complex.I]
|
||||
rw [partialDeriv_smul'₂]
|
||||
simp
|
||||
|
||||
have f_is_real_C2 : ContDiffOn ℝ 2 f t :=
|
||||
ContDiffOn.restrict_scalars ℝ (DifferentiableOn.contDiffOn h'f ht)
|
||||
rw [partialDeriv_commOn ht f_is_real_C2 Complex.I 1 w hw]
|
||||
rw [partialDeriv_eventuallyEq ℝ (CauchyRiemann'₆ (h'f w hw)) 1]
|
||||
rw [partialDeriv_commAt (HolomorphicAt_contDiffAt hw) Complex.I 1]
|
||||
rw [partialDeriv_eventuallyEq ℝ (CauchyRiemann'₆ hw) 1]
|
||||
rw [partialDeriv_smul'₂]
|
||||
simp
|
||||
|
||||
|
|
|
@ -1,9 +1,24 @@
|
|||
import Mathlib.Analysis.Complex.Basic
|
||||
import Nevanlinna.partialDeriv
|
||||
import Mathlib.Analysis.Complex.TaylorSeries
|
||||
import Mathlib.Analysis.Calculus.LineDeriv.Basic
|
||||
import Mathlib.Analysis.Calculus.ContDiff.Defs
|
||||
import Mathlib.Analysis.Calculus.FDeriv.Basic
|
||||
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
|
||||
import Mathlib.Analysis.RCLike.Basic
|
||||
import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv
|
||||
import Mathlib.Data.Complex.Module
|
||||
import Mathlib.Data.Complex.Order
|
||||
import Mathlib.Data.Complex.Exponential
|
||||
import Mathlib.Data.Fin.Tuple.Basic
|
||||
import Mathlib.Topology.Algebra.InfiniteSum.Module
|
||||
import Mathlib.Topology.Defs.Filter
|
||||
import Mathlib.Topology.Instances.RealVectorSpace
|
||||
import Nevanlinna.cauchyRiemann
|
||||
import Nevanlinna.laplace
|
||||
import Nevanlinna.complexHarmonic
|
||||
|
||||
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E]
|
||||
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F]
|
||||
|
||||
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F] [CompleteSpace F]
|
||||
|
||||
def HolomorphicAt (f : E → F) (x : E) : Prop :=
|
||||
∃ s ∈ nhds x, ∀ z ∈ s, DifferentiableAt ℂ f z
|
||||
|
@ -32,8 +47,17 @@ theorem HolomorphicAt_iff
|
|||
· assumption
|
||||
|
||||
|
||||
theorem HolomorphicAt_isOpen'
|
||||
{f : E → F} :
|
||||
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) :
|
||||
IsOpen { x : E | HolomorphicAt f x } := by
|
||||
|
||||
rw [← subset_interior_iff_isOpen]
|
||||
|
@ -53,37 +77,26 @@ theorem HolomorphicAt_isOpen'
|
|||
· 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 HolomorphicAt_contDiffAt
|
||||
{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
|
||||
|
||||
|
||||
theorem CauchyRiemann'₅
|
||||
{f : ℂ → F}
|
||||
|
|
Loading…
Reference in New Issue