nevanlinna/Nevanlinna/holomorphic.examples.lean

171 lines
4.7 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv
import Nevanlinna.complexHarmonic
import Nevanlinna.complexHarmonic
import Nevanlinna.holomorphicAt
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
have reg₀ : Differentiable g := by
let smulICLM : ≃L[] :=
{
toFun := fun x ↦ Complex.I • x
map_add' := fun x y => DistribSMul.smul_add Complex.I x y
map_smul' := fun m x => (smul_comm ((RingHom.id ) m) Complex.I x).symm
invFun := fun x ↦ (Complex.I)⁻¹ • x
left_inv := by
intro x
simp
rw [← mul_assoc, mul_comm]
simp
right_inv := by
intro x
simp
rw [← mul_assoc]
simp
continuous_toFun := continuous_const_smul Complex.I
continuous_invFun := continuous_const_smul (Complex.I)⁻¹
}
apply Differentiable.sub
apply Differentiable.comp
exact ContinuousLinearMap.differentiable Complex.ofRealCLM
intro z
sorry
apply Differentiable.comp
sorry
sorry
have reg₁ : Differentiable g := by
intro z
apply CauchyRiemann₇.2
constructor
· exact reg₀ z
· dsimp [g]
have : f_1 - Complex.I • f_I = f_1 + (- Complex.I • f_I) := by
rw [sub_eq_add_neg]
simp
rw [this, partialDeriv_add₂, partialDeriv_add₂]
simp
dsimp [f_1, f_I]
sorry
sorry
sorry
sorry
sorry
sorry