nevanlinna/Nevanlinna/holomorphic_examples.lean

315 lines
9.8 KiB
Plaintext
Raw Normal View History

2024-07-29 15:50:51 +02:00
import Nevanlinna.complexHarmonic
import Nevanlinna.holomorphicAt
2024-08-08 12:36:06 +02:00
import Nevanlinna.holomorphic_primitive2
2024-07-31 09:40:35 +02:00
import Nevanlinna.mathlibAddOn
2024-07-29 15:50:51 +02:00
theorem CauchyRiemann₆
2024-07-31 09:40:35 +02:00
{E : Type*} [NormedAddCommGroup E] [NormedSpace E]
{F : Type*} [NormedAddCommGroup F] [NormedSpace F]
2024-07-29 15:50:51 +02:00
{f : E → F}
2024-07-31 09:40:35 +02:00
{z : E} :
2024-07-29 15:50:51 +02:00
(DifferentiableAt f z) ↔ (DifferentiableAt f z) ∧ ∀ e, partialDeriv (Complex.I • e) f z = Complex.I • partialDeriv e f z := by
constructor
2024-07-31 09:40:35 +02:00
· -- Direction "→"
2024-07-29 15:50:51 +02:00
intro h
constructor
· exact DifferentiableAt.restrictScalars h
· unfold partialDeriv
conv =>
intro e
left
rw [DifferentiableAt.fderiv_restrictScalars h]
simp
rw [← mul_one Complex.I]
rw [← smul_eq_mul]
conv =>
intro e
right
right
rw [DifferentiableAt.fderiv_restrictScalars h]
simp
2024-07-31 09:40:35 +02:00
· -- Direction "←"
2024-07-29 15:50:51 +02:00
intro ⟨h₁, h₂⟩
apply (differentiableAt_iff_restrictScalars h₁).2
use {
toFun := fderiv f z
map_add' := fun x y => ContinuousLinearMap.map_add (fderiv f z) x y
map_smul' := by
simp
intro m x
have : m = m.re + m.im • Complex.I := by simp
rw [this, add_smul, add_smul, ContinuousLinearMap.map_add]
congr
simp
rw [smul_assoc, smul_assoc, ContinuousLinearMap.map_smul (fderiv f z) m.2]
congr
exact h₂ x
}
rfl
theorem CauchyRiemann₇
2024-07-31 09:40:35 +02:00
{F : Type*} [NormedAddCommGroup F] [NormedSpace F]
2024-07-29 15:50:51 +02:00
{f : → F}
2024-07-31 09:40:35 +02:00
{z : } :
2024-07-29 15:50:51 +02:00
(DifferentiableAt f z) ↔ (DifferentiableAt f z) ∧ partialDeriv Complex.I f z = Complex.I • partialDeriv 1 f z := by
constructor
· intro hf
constructor
· exact (CauchyRiemann₆.1 hf).1
· let A := (CauchyRiemann₆.1 hf).2 1
simp at A
exact A
· intro ⟨h₁, h₂⟩
apply CauchyRiemann₆.2
constructor
· exact h₁
· intro e
2024-07-31 09:40:35 +02:00
have : Complex.I • e = e • Complex.I := by
2024-07-29 15:50:51 +02:00
rw [smul_eq_mul, smul_eq_mul]
exact CommMonoid.mul_comm Complex.I e
rw [this]
have : e = e.re + e.im • Complex.I := by simp
rw [this, add_smul, partialDeriv_add₁, partialDeriv_add₁]
simp
rw [← smul_eq_mul]
have : partialDeriv ((e.re : ) • Complex.I) f = partialDeriv ((e.re : ) • Complex.I) f := by rfl
rw [← this, partialDeriv_smul₁ ]
have : (e.re : ) = (e.re : ) • (1 : ) := by simp
rw [this, partialDeriv_smul₁ ]
have : partialDeriv ((e.im : ) * Complex.I) f = partialDeriv ((e.im : ) • Complex.I) f := by rfl
rw [this, partialDeriv_smul₁ ]
simp
rw [h₂]
rw [smul_comm]
congr
rw [mul_assoc]
simp
nth_rw 2 [smul_comm]
rw [← smul_assoc]
simp
have : - (e.im : ) = (-e.im : ) • (1 : ) := by simp
rw [this, partialDeriv_smul₁ ]
simp
2024-07-31 09:40:35 +02:00
2024-07-29 15:50:51 +02:00
/-
2024-07-31 09:40:35 +02:00
A harmonic, real-valued function on is the real part of a suitable holomorphic
function.
2024-07-29 15:50:51 +02:00
-/
theorem harmonic_is_realOfHolomorphic
{f : }
2024-08-08 12:36:06 +02:00
{z : }
{R : }
(hR : 0 < R)
(hf : ∀ x ∈ Metric.ball z R, HarmonicAt f x) :
2024-08-08 13:30:59 +02:00
∃ F : , (∀ x ∈ Metric.ball z R, HolomorphicAt F x) ∧ (Set.EqOn (Complex.reCLM ∘ F) f (Metric.ball z R)) := by
2024-07-29 15:50:51 +02:00
let f_1 : := Complex.ofRealCLM ∘ (partialDeriv 1 f)
let f_I : := Complex.ofRealCLM ∘ (partialDeriv Complex.I f)
let g : := f_1 - Complex.I • f_I
2024-08-08 12:36:06 +02:00
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
2024-07-29 15:50:51 +02:00
2024-08-08 12:36:06 +02:00
have reg₂f : ContDiffOn 2 f (Metric.ball z R) := by
apply contDiffOn_if_contDiffAt
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
2024-07-30 10:52:12 +02:00
dsimp [f_1]
apply ContDiffAt.continuousLinearMap_comp
2024-08-08 12:36:06 +02:00
exact partialDeriv_contDiffAt (hf z hz).1 1
2024-07-29 15:50:51 +02:00
2024-08-08 12:36:06 +02:00
have reg₁f_I : ContDiffOn 1 f_I (Metric.ball z R) := by
apply contDiffOn_if_contDiffAt'
intro z hz
2024-07-30 10:52:12 +02:00
dsimp [f_I]
apply ContDiffAt.continuousLinearMap_comp
2024-08-08 12:36:06 +02:00
exact partialDeriv_contDiffAt (hf z hz).1 Complex.I
2024-07-30 10:52:12 +02:00
2024-08-08 12:36:06 +02:00
have reg₁g : ContDiffOn 1 g (Metric.ball z R) := by
2024-07-30 10:52:12 +02:00
dsimp [g]
2024-08-08 12:36:06 +02:00
apply ContDiffOn.sub
2024-07-30 10:52:12 +02:00
exact reg₁f_1
2024-08-08 12:36:06 +02:00
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := by rfl
rw [this]
apply ContDiffOn.const_smul
2024-07-30 10:52:12 +02:00
exact reg₁f_I
2024-07-31 09:40:35 +02:00
2024-08-08 12:36:06 +02:00
have reg₁ : DifferentiableOn g (Metric.ball z R) := by
intro x hx
apply DifferentiableAt.differentiableWithinAt
2024-07-29 15:50:51 +02:00
apply CauchyRiemann₇.2
constructor
2024-08-08 12:36:06 +02:00
· apply DifferentiableWithinAt.differentiableAt (reg₁g.differentiableOn le_rfl x hx)
apply IsOpen.mem_nhds Metric.isOpen_ball hx
2024-07-29 15:50:51 +02:00
· dsimp [g]
2024-08-08 12:36:06 +02:00
rw [partialDeriv_sub₂_differentiableAt, partialDeriv_sub₂_differentiableAt]
2024-07-29 15:50:51 +02:00
dsimp [f_1, f_I]
2024-07-30 10:52:12 +02:00
rw [partialDeriv_smul'₂, partialDeriv_smul'₂]
2024-08-08 12:36:06 +02:00
rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt]
simp
rw [partialDeriv_compContLinAt, partialDeriv_compContLinAt]
2024-07-30 10:52:12 +02:00
rw [mul_sub]
simp
rw [← mul_assoc]
simp
rw [add_comm, sub_eq_add_neg]
congr 1
2024-08-08 12:36:06 +02:00
· rw [partialDeriv_commOn _ reg₂f Complex.I 1]
exact hx
exact Metric.isOpen_ball
· let A := Filter.EventuallyEq.eq_of_nhds (hf x hx).2
2024-07-30 10:52:12 +02:00
simp at A
unfold Complex.laplace at A
conv =>
right
right
2024-08-08 12:36:06 +02:00
rw [← sub_zero (partialDeriv 1 (partialDeriv 1 f) x)]
2024-07-30 10:52:12 +02:00
rw [← A]
simp
2024-07-29 15:50:51 +02:00
2024-08-08 12:36:06 +02:00
--DifferentiableAt (partialDeriv _ f)
2024-07-30 10:52:12 +02:00
repeat
2024-08-08 12:36:06 +02:00
apply ContDiffAt.differentiableAt
apply partialDeriv_contDiffAt (hf x hx).1
2024-07-30 10:52:12 +02:00
apply le_rfl
2024-08-08 12:36:06 +02:00
-- DifferentiableAt f_1 x
apply (reg₁f_1.differentiableOn le_rfl).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
-- DifferentiableAt (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
2024-07-30 10:52:12 +02:00
-- Differentiable f_1
2024-08-08 12:36:06 +02:00
apply (reg₁f_1.differentiableOn le_rfl).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
2024-07-30 10:52:12 +02:00
-- Differentiable (Complex.I • f_I)
2024-08-08 12:36:06 +02:00
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
2024-07-31 09:40:35 +02:00
2024-08-08 12:36:06 +02:00
let F := fun z' ↦ (primitive z g) z' + f z
2024-07-30 16:45:26 +02:00
2024-08-08 12:36:06 +02:00
have regF : DifferentiableOn F (Metric.ball z R) := by
apply DifferentiableOn.add
apply primitive_differentiableOn reg₁
2024-07-30 15:11:57 +02:00
simp
2024-07-30 14:11:39 +02:00
2024-08-08 12:36:06 +02:00
have pF'' : ∀ x ∈ Metric.ball z R, (fderiv F x) = ContinuousLinearMap.lsmul (g x) := by
intro x hx
have : DifferentiableAt F x := by
apply (regF x hx).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
rw [DifferentiableAt.fderiv_restrictScalars this]
2024-07-30 15:11:57 +02:00
dsimp [F]
rw [fderiv_add_const]
2024-08-08 12:36:06 +02:00
rw [primitive_fderiv' reg₁ x hx]
2024-07-30 15:11:57 +02:00
exact rfl
2024-07-30 14:11:39 +02:00
2024-07-30 12:19:59 +02:00
use F
2024-07-30 12:19:59 +02:00
constructor
2024-08-08 13:30:59 +02:00
· -- ∀ x ∈ Metric.ball z R, HolomorphicAt F x
2024-08-08 12:40:32 +02:00
intro x hx
2024-07-30 12:19:59 +02:00
apply HolomorphicAt_iff.2
2024-08-08 12:40:32 +02:00
use Metric.ball z R
2024-07-30 12:19:59 +02:00
constructor
2024-08-08 12:40:32 +02:00
· exact Metric.isOpen_ball
2024-07-30 12:19:59 +02:00
· constructor
2024-08-08 12:40:32 +02:00
· assumption
· intro w hw
apply (regF w hw).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hw
2024-08-08 13:30:59 +02:00
· -- Set.EqOn (⇑Complex.reCLM ∘ F) f (Metric.ball z R)
have : DifferentiableOn (Complex.reCLM ∘ F) (Metric.ball z R) := by
apply DifferentiableOn.comp
apply Differentiable.differentiableOn
apply ContinuousLinearMap.differentiable Complex.reCLM
apply regF.restrictScalars
exact Set.mapsTo'.mpr fun ⦃a⦄ _ => hR
have hz : z ∈ Metric.ball z R := by exact Metric.mem_ball_self hR
apply Convex.eqOn_of_fderivWithin_eq _ this _ _ _ hz _
exact convex_ball z R
apply reg₂f.differentiableOn one_le_two
apply IsOpen.uniqueDiffOn Metric.isOpen_ball
2024-07-30 16:45:26 +02:00
2024-08-08 13:30:59 +02:00
intro x hx
rw [fderivWithin_eq_fderiv, fderivWithin_eq_fderiv]
2024-07-30 14:11:39 +02:00
rw [fderiv.comp]
simp
apply ContinuousLinearMap.ext
intro w
simp
2024-07-30 15:11:57 +02:00
rw [pF'']
2024-07-30 14:11:39 +02:00
dsimp [g, f_1, f_I, partialDeriv]
simp
have : w = w.re • 1 + w.im • Complex.I := by simp
nth_rw 3 [this]
rw [(fderiv f x).map_add]
rw [(fderiv f x).map_smul, (fderiv f x).map_smul]
rw [smul_eq_mul, smul_eq_mul]
ring
2024-08-08 13:30:59 +02:00
assumption
exact ContinuousLinearMap.differentiableAt Complex.reCLM
apply (regF.restrictScalars x hx).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
apply IsOpen.uniqueDiffOn Metric.isOpen_ball
assumption
apply (reg₂f.differentiableOn one_le_two).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
apply IsOpen.uniqueDiffOn Metric.isOpen_ball
assumption
-- DifferentiableAt (⇑Complex.reCLM ∘ F) x
apply DifferentiableAt.comp
apply Differentiable.differentiableAt
exact ContinuousLinearMap.differentiable Complex.reCLM
apply (regF.restrictScalars x hx).differentiableAt
apply IsOpen.mem_nhds Metric.isOpen_ball hx
--
dsimp [F]
rw [primitive_zeroAtBasepoint]
simp