nevanlinna/Nevanlinna/holomorphic_examples.lean

214 lines
6.3 KiB
Plaintext
Raw Normal View History

2024-07-29 15:50:51 +02:00
import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv
import Nevanlinna.complexHarmonic
import Nevanlinna.holomorphicAt
2024-07-30 12:19:59 +02:00
import Nevanlinna.holomorphic_primitive
2024-07-29 15:50:51 +02:00
theorem CauchyRiemann₆
{E : Type*} [NormedAddCommGroup E] [NormedSpace E]
{F : Type*} [NormedAddCommGroup F] [NormedSpace F]
{f : E → F}
{z : E} :
(DifferentiableAt f z) ↔ (DifferentiableAt f z) ∧ ∀ e, partialDeriv (Complex.I • e) f z = Complex.I • partialDeriv e f z := by
constructor
· -- Direction "→"
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
· -- Direction "←"
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₇
{F : Type*} [NormedAddCommGroup F] [NormedSpace F]
{f : → F}
{z : } :
(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
have : Complex.I • e = e • Complex.I := by
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
/-
A harmonic, real-valued function on is the real part of a suitable holomorphic function.
-/
theorem harmonic_is_realOfHolomorphic
{f : }
(hf : ∀ z, HarmonicAt f z) :
∃ F : , ∀ z, (HolomorphicAt F z ∧ ((F z).re = f z)) := by
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-07-30 10:52:12 +02:00
have reg₂f : ContDiff 2 f := by
apply contDiff_iff_contDiffAt.mpr
intro z
exact (hf z).1
2024-07-29 15:50:51 +02:00
2024-07-30 10:52:12 +02:00
have reg₁f_1 : ContDiff 1 f_1 := by
apply contDiff_iff_contDiffAt.mpr
2024-07-29 15:50:51 +02:00
intro z
2024-07-30 10:52:12 +02:00
dsimp [f_1]
apply ContDiffAt.continuousLinearMap_comp
exact partialDeriv_contDiffAt (hf z).1 1
2024-07-29 15:50:51 +02:00
2024-07-30 10:52:12 +02:00
have reg₁f_I : ContDiff 1 f_I := by
apply contDiff_iff_contDiffAt.mpr
intro z
dsimp [f_I]
apply ContDiffAt.continuousLinearMap_comp
exact partialDeriv_contDiffAt (hf z).1 Complex.I
have reg₁g : ContDiff 1 g := by
dsimp [g]
apply ContDiff.sub
exact reg₁f_1
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := rfl
rw [this]
apply ContDiff.const_smul
exact reg₁f_I
2024-07-29 15:50:51 +02:00
have reg₁ : Differentiable g := by
intro z
apply CauchyRiemann₇.2
constructor
2024-07-30 10:52:12 +02:00
· apply Differentiable.differentiableAt
apply ContDiff.differentiable
exact reg₁g
rfl
2024-07-29 15:50:51 +02:00
· dsimp [g]
2024-07-30 10:52:12 +02:00
rw [partialDeriv_sub₂, partialDeriv_sub₂]
2024-07-29 15:50:51 +02:00
simp
dsimp [f_1, f_I]
2024-07-30 10:52:12 +02:00
rw [partialDeriv_smul'₂, partialDeriv_smul'₂]
rw [partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin, partialDeriv_compContLin]
rw [mul_sub]
simp
rw [← mul_assoc]
simp
rw [add_comm, sub_eq_add_neg]
congr 1
· rw [partialDeriv_comm reg₂f Complex.I 1]
· let A := Filter.EventuallyEq.eq_of_nhds (hf z).2
simp at A
unfold Complex.laplace at A
conv =>
right
right
rw [← sub_zero (partialDeriv 1 (partialDeriv 1 f) z)]
rw [← A]
simp
2024-07-29 15:50:51 +02:00
2024-07-30 10:52:12 +02:00
--Differentiable (partialDeriv _ f)
repeat
apply ContDiff.differentiable
apply contDiff_iff_contDiffAt.mpr
exact fun w ↦ partialDeriv_contDiffAt (hf w).1 _
apply le_rfl
-- Differentiable f_1
exact reg₁f_1.differentiable le_rfl
-- Differentiable (Complex.I • f_I)
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := rfl
rw [this]
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)
have : Complex.I • f_I = fun x ↦ Complex.I • f_I x := rfl
rw [this]
apply Differentiable.const_smul
exact reg₁f_I.differentiable le_rfl
2024-07-30 12:19:59 +02:00
let F := primitive 0 g
use F
intro z
constructor
· -- HolomorphicAt F z
apply HolomorphicAt_iff.2
use {z : | true}
constructor
· exact isOpen_const
· constructor
· simp
· intro w hw
let A : HasDerivAt (primitive 0 g) (g w) w := primitive_fderiv g reg₁
apply A.differentiableAt
· -- (F z).re = f z
sorry