working…
This commit is contained in:
parent
acb1f34879
commit
3e924d5b4a
|
@ -1,6 +1,6 @@
|
||||||
import Nevanlinna.complexHarmonic
|
import Nevanlinna.complexHarmonic
|
||||||
import Nevanlinna.holomorphicAt
|
import Nevanlinna.holomorphicAt
|
||||||
import Nevanlinna.holomorphic_primitive
|
import Nevanlinna.holomorphic_primitive2
|
||||||
import Nevanlinna.mathlibAddOn
|
import Nevanlinna.mathlibAddOn
|
||||||
|
|
||||||
|
|
||||||
|
@ -108,103 +108,145 @@ function.
|
||||||
|
|
||||||
theorem harmonic_is_realOfHolomorphic
|
theorem harmonic_is_realOfHolomorphic
|
||||||
{f : ℂ → ℝ}
|
{f : ℂ → ℝ}
|
||||||
(hf : ∀ z, HarmonicAt f z) :
|
{z : ℂ}
|
||||||
∃ F : ℂ → ℂ, (∀ z, HolomorphicAt F z) ∧ (Complex.reCLM ∘ F = f) := by
|
{R : ℝ}
|
||||||
|
(hR : 0 < R)
|
||||||
|
(hf : ∀ x ∈ Metric.ball z R, HarmonicAt f x) :
|
||||||
|
∃ F : ℂ → ℂ, (∀ z ∈ Metric.ball z R, HolomorphicAt F z) ∧ (Set.EqOn (Complex.reCLM ∘ F) f (Metric.ball z R)) := by
|
||||||
|
|
||||||
let f_1 : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ 1 f)
|
let f_1 : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ 1 f)
|
||||||
let f_I : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ Complex.I f)
|
let f_I : ℂ → ℂ := Complex.ofRealCLM ∘ (partialDeriv ℝ Complex.I f)
|
||||||
|
|
||||||
let g : ℂ → ℂ := f_1 - Complex.I • f_I
|
let g : ℂ → ℂ := f_1 - Complex.I • f_I
|
||||||
|
|
||||||
have reg₂f : ContDiff ℝ 2 f := by
|
have contDiffOn_if_contDiffAt
|
||||||
apply contDiff_iff_contDiffAt.mpr
|
{f' : ℂ → ℝ}
|
||||||
intro z
|
{z' : ℂ}
|
||||||
exact (hf z).1
|
{R' : ℝ}
|
||||||
|
{n' : ℕ}
|
||||||
|
(hf' : ∀ x ∈ Metric.ball z' R', ContDiffAt ℝ n' f' x) :
|
||||||
|
ContDiffOn ℝ n' f' (Metric.ball z' R') := by
|
||||||
|
intro z hz
|
||||||
|
apply ContDiffAt.contDiffWithinAt
|
||||||
|
exact hf' z hz
|
||||||
|
|
||||||
have reg₁f_1 : ContDiff ℝ 1 f_1 := by
|
have reg₂f : ContDiffOn ℝ 2 f (Metric.ball z R) := by
|
||||||
apply contDiff_iff_contDiffAt.mpr
|
apply contDiffOn_if_contDiffAt
|
||||||
intro z
|
intro x hx
|
||||||
|
exact (hf x hx).1
|
||||||
|
|
||||||
|
have contDiffOn_if_contDiffAt'
|
||||||
|
{f' : ℂ → ℂ}
|
||||||
|
{z' : ℂ}
|
||||||
|
{R' : ℝ}
|
||||||
|
{n' : ℕ}
|
||||||
|
(hf' : ∀ x ∈ Metric.ball z' R', ContDiffAt ℝ n' f' x) :
|
||||||
|
ContDiffOn ℝ n' f' (Metric.ball z' R') := by
|
||||||
|
intro z hz
|
||||||
|
apply ContDiffAt.contDiffWithinAt
|
||||||
|
exact hf' z hz
|
||||||
|
|
||||||
|
have reg₁f_1 : ContDiffOn ℝ 1 f_1 (Metric.ball z R) := by
|
||||||
|
apply contDiffOn_if_contDiffAt'
|
||||||
|
intro z hz
|
||||||
dsimp [f_1]
|
dsimp [f_1]
|
||||||
apply ContDiffAt.continuousLinearMap_comp
|
apply ContDiffAt.continuousLinearMap_comp
|
||||||
exact partialDeriv_contDiffAt ℝ (hf z).1 1
|
exact partialDeriv_contDiffAt ℝ (hf z hz).1 1
|
||||||
|
|
||||||
have reg₁f_I : ContDiff ℝ 1 f_I := by
|
have reg₁f_I : ContDiffOn ℝ 1 f_I (Metric.ball z R) := by
|
||||||
apply contDiff_iff_contDiffAt.mpr
|
apply contDiffOn_if_contDiffAt'
|
||||||
intro z
|
intro z hz
|
||||||
dsimp [f_I]
|
dsimp [f_I]
|
||||||
apply ContDiffAt.continuousLinearMap_comp
|
apply ContDiffAt.continuousLinearMap_comp
|
||||||
exact partialDeriv_contDiffAt ℝ (hf z).1 Complex.I
|
exact partialDeriv_contDiffAt ℝ (hf z hz).1 Complex.I
|
||||||
|
|
||||||
have reg₁g : ContDiff ℝ 1 g := by
|
have reg₁g : ContDiffOn ℝ 1 g (Metric.ball z R) := by
|
||||||
dsimp [g]
|
dsimp [g]
|
||||||
apply ContDiff.sub
|
apply ContDiffOn.sub
|
||||||
exact reg₁f_1
|
exact reg₁f_1
|
||||||
apply ContDiff.const_smul'
|
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl
|
||||||
|
rw [this]
|
||||||
|
apply ContDiffOn.const_smul
|
||||||
exact reg₁f_I
|
exact reg₁f_I
|
||||||
|
|
||||||
have reg₁ : Differentiable ℂ g := by
|
have reg₁ : DifferentiableOn ℂ g (Metric.ball z R) := by
|
||||||
intro z
|
intro x hx
|
||||||
|
apply DifferentiableAt.differentiableWithinAt
|
||||||
apply CauchyRiemann₇.2
|
apply CauchyRiemann₇.2
|
||||||
constructor
|
constructor
|
||||||
· apply Differentiable.differentiableAt
|
· apply DifferentiableWithinAt.differentiableAt (reg₁g.differentiableOn le_rfl x hx)
|
||||||
apply ContDiff.differentiable
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
exact reg₁g
|
|
||||||
rfl
|
|
||||||
· dsimp [g]
|
· dsimp [g]
|
||||||
rw [partialDeriv_sub₂, partialDeriv_sub₂]
|
rw [partialDeriv_sub₂_differentiableAt, partialDeriv_sub₂_differentiableAt]
|
||||||
simp
|
|
||||||
dsimp [f_1, f_I]
|
dsimp [f_1, f_I]
|
||||||
rw [partialDeriv_smul'₂, partialDeriv_smul'₂]
|
rw [partialDeriv_smul'₂, partialDeriv_smul'₂]
|
||||||
rw [partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin]
|
rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt]
|
||||||
|
simp
|
||||||
|
rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt]
|
||||||
rw [mul_sub]
|
rw [mul_sub]
|
||||||
simp
|
simp
|
||||||
rw [← mul_assoc]
|
rw [← mul_assoc]
|
||||||
simp
|
simp
|
||||||
rw [add_comm, sub_eq_add_neg]
|
rw [add_comm, sub_eq_add_neg]
|
||||||
congr 1
|
congr 1
|
||||||
· rw [partialDeriv_comm reg₂f Complex.I 1]
|
· rw [partialDeriv_commOn _ reg₂f Complex.I 1]
|
||||||
· let A := Filter.EventuallyEq.eq_of_nhds (hf z).2
|
exact hx
|
||||||
|
exact Metric.isOpen_ball
|
||||||
|
· let A := Filter.EventuallyEq.eq_of_nhds (hf x hx).2
|
||||||
simp at A
|
simp at A
|
||||||
unfold Complex.laplace at A
|
unfold Complex.laplace at A
|
||||||
conv =>
|
conv =>
|
||||||
right
|
right
|
||||||
right
|
right
|
||||||
rw [← sub_zero (partialDeriv ℝ 1 (partialDeriv ℝ 1 f) z)]
|
rw [← sub_zero (partialDeriv ℝ 1 (partialDeriv ℝ 1 f) x)]
|
||||||
rw [← A]
|
rw [← A]
|
||||||
simp
|
simp
|
||||||
|
|
||||||
--Differentiable ℝ (partialDeriv ℝ _ f)
|
--DifferentiableAt ℝ (partialDeriv ℝ _ f)
|
||||||
repeat
|
repeat
|
||||||
apply ContDiff.differentiable
|
apply ContDiffAt.differentiableAt
|
||||||
apply contDiff_iff_contDiffAt.mpr
|
apply partialDeriv_contDiffAt ℝ (hf x hx).1
|
||||||
exact fun w ↦ partialDeriv_contDiffAt ℝ (hf w).1 _
|
|
||||||
apply le_rfl
|
apply le_rfl
|
||||||
-- Differentiable ℝ f_1
|
|
||||||
exact reg₁f_1.differentiable le_rfl
|
|
||||||
-- Differentiable ℝ (Complex.I • f_I)
|
|
||||||
apply Differentiable.const_smul'
|
|
||||||
exact reg₁f_I.differentiable le_rfl
|
|
||||||
-- Differentiable ℝ f_1
|
|
||||||
exact reg₁f_1.differentiable le_rfl
|
|
||||||
-- Differentiable ℝ (Complex.I • f_I)
|
|
||||||
apply Differentiable.const_smul'
|
|
||||||
exact reg₁f_I.differentiable le_rfl
|
|
||||||
|
|
||||||
let F := fun z ↦ (primitive 0 g) z + f 0
|
-- DifferentiableAt ℝ f_1 x
|
||||||
|
apply (reg₁f_1.differentiableOn le_rfl).differentiableAt
|
||||||
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
|
|
||||||
have regF : Differentiable ℂ F := by
|
-- DifferentiableAt ℝ (Complex.I • f_I)
|
||||||
apply Differentiable.add
|
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl
|
||||||
apply primitive_differentiable reg₁
|
rw [this]
|
||||||
|
apply DifferentiableAt.const_smul
|
||||||
|
apply (reg₁f_I.differentiableOn le_rfl).differentiableAt
|
||||||
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
|
|
||||||
|
-- Differentiable ℝ f_1
|
||||||
|
apply (reg₁f_1.differentiableOn le_rfl).differentiableAt
|
||||||
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
|
|
||||||
|
-- Differentiable ℝ (Complex.I • f_I)
|
||||||
|
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl
|
||||||
|
rw [this]
|
||||||
|
apply DifferentiableAt.const_smul
|
||||||
|
apply (reg₁f_I.differentiableOn le_rfl).differentiableAt
|
||||||
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
|
|
||||||
|
let F := fun z' ↦ (primitive z g) z' + f z
|
||||||
|
|
||||||
|
have regF : DifferentiableOn ℂ F (Metric.ball z R) := by
|
||||||
|
apply DifferentiableOn.add
|
||||||
|
apply primitive_differentiableOn reg₁
|
||||||
simp
|
simp
|
||||||
|
|
||||||
have pF'' : ∀ x, (fderiv ℝ F x) = ContinuousLinearMap.lsmul ℝ ℂ (g x) := by
|
have pF'' : ∀ x ∈ Metric.ball z R, (fderiv ℝ F x) = ContinuousLinearMap.lsmul ℝ ℂ (g x) := by
|
||||||
intro x
|
intro x hx
|
||||||
rw [DifferentiableAt.fderiv_restrictScalars ℝ (regF x)]
|
have : DifferentiableAt ℂ F x := by
|
||||||
|
apply (regF x hx).differentiableAt
|
||||||
|
apply IsOpen.mem_nhds Metric.isOpen_ball hx
|
||||||
|
rw [DifferentiableAt.fderiv_restrictScalars ℝ this]
|
||||||
dsimp [F]
|
dsimp [F]
|
||||||
rw [fderiv_add_const]
|
rw [fderiv_add_const]
|
||||||
rw [primitive_fderiv']
|
rw [primitive_fderiv' reg₁ x hx]
|
||||||
exact rfl
|
exact rfl
|
||||||
exact reg₁
|
|
||||||
|
|
||||||
use F
|
use F
|
||||||
|
|
||||||
|
|
|
@ -781,12 +781,11 @@ theorem primitive_hasDerivAt
|
||||||
apply hasDerivAt_const
|
apply hasDerivAt_const
|
||||||
|
|
||||||
|
|
||||||
|
theorem primitive_differentiableOn
|
||||||
theorem primitive_differentiable
|
|
||||||
{E : Type u} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E]
|
{E : Type u} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E]
|
||||||
{f : ℂ → E}
|
{f : ℂ → E}
|
||||||
(z₀ : ℂ)
|
{z₀ : ℂ}
|
||||||
(R : ℝ)
|
{R : ℝ}
|
||||||
(hf : DifferentiableOn ℂ f (Metric.ball z₀ R))
|
(hf : DifferentiableOn ℂ f (Metric.ball z₀ R))
|
||||||
:
|
:
|
||||||
DifferentiableOn ℂ (primitive z₀ f) (Metric.ball z₀ R) := by
|
DifferentiableOn ℂ (primitive z₀ f) (Metric.ball z₀ R) := by
|
||||||
|
|
|
@ -102,20 +102,6 @@ theorem partialDeriv_add₂ {f₁ f₂ : E → F} (h₁ : Differentiable 𝕜 f
|
||||||
simp
|
simp
|
||||||
|
|
||||||
|
|
||||||
theorem partialDeriv_sub₂ {f₁ f₂ : E → F} (h₁ : Differentiable 𝕜 f₁) (h₂ : Differentiable 𝕜 f₂) : ∀ v : E, partialDeriv 𝕜 v (f₁ - f₂) = (partialDeriv 𝕜 v f₁) - (partialDeriv 𝕜 v f₂) := by
|
|
||||||
unfold partialDeriv
|
|
||||||
intro v
|
|
||||||
have : f₁ - f₂ = fun y ↦ f₁ y - f₂ y := by rfl
|
|
||||||
rw [this]
|
|
||||||
conv =>
|
|
||||||
left
|
|
||||||
intro w
|
|
||||||
left
|
|
||||||
rw [fderiv_sub (h₁ w) (h₂ w)]
|
|
||||||
funext w
|
|
||||||
simp
|
|
||||||
|
|
||||||
|
|
||||||
theorem partialDeriv_add₂_differentiableAt
|
theorem partialDeriv_add₂_differentiableAt
|
||||||
{f₁ f₂ : E → F}
|
{f₁ f₂ : E → F}
|
||||||
{v : E}
|
{v : E}
|
||||||
|
@ -153,6 +139,57 @@ theorem partialDeriv_add₂_contDiffAt
|
||||||
exact (hf₂' x (Set.mem_of_mem_inter_right hx)).differentiableAt
|
exact (hf₂' x (Set.mem_of_mem_inter_right hx)).differentiableAt
|
||||||
|
|
||||||
|
|
||||||
|
theorem partialDeriv_sub₂ {f₁ f₂ : E → F} (h₁ : Differentiable 𝕜 f₁) (h₂ : Differentiable 𝕜 f₂) : ∀ v : E, partialDeriv 𝕜 v (f₁ - f₂) = (partialDeriv 𝕜 v f₁) - (partialDeriv 𝕜 v f₂) := by
|
||||||
|
unfold partialDeriv
|
||||||
|
intro v
|
||||||
|
have : f₁ - f₂ = fun y ↦ f₁ y - f₂ y := by rfl
|
||||||
|
rw [this]
|
||||||
|
conv =>
|
||||||
|
left
|
||||||
|
intro w
|
||||||
|
left
|
||||||
|
rw [fderiv_sub (h₁ w) (h₂ w)]
|
||||||
|
funext w
|
||||||
|
simp
|
||||||
|
|
||||||
|
|
||||||
|
theorem partialDeriv_sub₂_differentiableAt
|
||||||
|
{f₁ f₂ : E → F}
|
||||||
|
{v : E}
|
||||||
|
{x : E}
|
||||||
|
(h₁ : DifferentiableAt 𝕜 f₁ x)
|
||||||
|
(h₂ : DifferentiableAt 𝕜 f₂ x) :
|
||||||
|
partialDeriv 𝕜 v (f₁ - f₂) x = (partialDeriv 𝕜 v f₁) x - (partialDeriv 𝕜 v f₂) x := by
|
||||||
|
|
||||||
|
unfold partialDeriv
|
||||||
|
have : f₁ - f₂ = fun y ↦ f₁ y - f₂ y := by rfl
|
||||||
|
rw [this]
|
||||||
|
rw [fderiv_sub h₁ h₂]
|
||||||
|
rfl
|
||||||
|
|
||||||
|
|
||||||
|
theorem partialDeriv_sub₂_contDiffAt
|
||||||
|
{f₁ f₂ : E → F}
|
||||||
|
{v : E}
|
||||||
|
{x : E}
|
||||||
|
(h₁ : ContDiffAt 𝕜 1 f₁ x)
|
||||||
|
(h₂ : ContDiffAt 𝕜 1 f₂ x) :
|
||||||
|
partialDeriv 𝕜 v (f₁ - f₂) =ᶠ[nhds x] (partialDeriv 𝕜 v f₁) - (partialDeriv 𝕜 v f₂) := by
|
||||||
|
|
||||||
|
obtain ⟨f₁', u₁, hu₁, _, hf₁'⟩ := contDiffAt_one_iff.1 h₁
|
||||||
|
obtain ⟨f₂', u₂, hu₂, _, hf₂'⟩ := contDiffAt_one_iff.1 h₂
|
||||||
|
|
||||||
|
apply Filter.eventuallyEq_iff_exists_mem.2
|
||||||
|
use u₁ ∩ u₂
|
||||||
|
constructor
|
||||||
|
· exact Filter.inter_mem hu₁ hu₂
|
||||||
|
· intro x hx
|
||||||
|
simp
|
||||||
|
apply partialDeriv_sub₂_differentiableAt 𝕜
|
||||||
|
exact (hf₁' x (Set.mem_of_mem_inter_left hx)).differentiableAt
|
||||||
|
exact (hf₂' x (Set.mem_of_mem_inter_right hx)).differentiableAt
|
||||||
|
|
||||||
|
|
||||||
theorem partialDeriv_compContLin
|
theorem partialDeriv_compContLin
|
||||||
{f : E → F}
|
{f : E → F}
|
||||||
{l : F →L[𝕜] G}
|
{l : F →L[𝕜] G}
|
||||||
|
|
Loading…
Reference in New Issue