This commit is contained in:
Stefan Kebekus 2024-05-06 09:01:43 +02:00
parent a55b084031
commit c44f4fe3b0
2 changed files with 6 additions and 3 deletions

View File

@ -3,6 +3,7 @@ import Mathlib.Analysis.Complex.RealDeriv
variable {z : } {f : }
theorem CauchyRiemann₁ : (DifferentiableAt f z)
→ (fderiv f z) Complex.I = Complex.I * (fderiv f z) 1 := by
intro h
@ -10,6 +11,7 @@ theorem CauchyRiemann₁ : (DifferentiableAt f z)
nth_rewrite 1 [← mul_one Complex.I]
exact ContinuousLinearMap.map_smul_of_tower (fderiv f z) Complex.I 1
theorem CauchyRiemann₂ : (DifferentiableAt f z)
→ lineDeriv f z Complex.I = Complex.I * lineDeriv f z 1 := by
intro h
@ -17,6 +19,7 @@ theorem CauchyRiemann₂ : (DifferentiableAt f z)
rw [DifferentiableAt.lineDeriv_eq_fderiv (h.restrictScalars )]
exact CauchyRiemann₁ h
theorem CauchyRiemann₃ : (DifferentiableAt f z)
→ (lineDeriv (Complex.reCLM ∘ f) z 1 = lineDeriv (Complex.imCLM ∘ f) z Complex.I)
∧ (lineDeriv (Complex.reCLM ∘ f) z Complex.I = -lineDeriv (Complex.imCLM ∘ f) z 1)

View File

@ -20,7 +20,6 @@ noncomputable def Complex.laplace : () → () := by
def Harmonic (f : ) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
#check second_derivative_symmetric
lemma zwoDiff (f : × ) (h : ContDiff 2 f) : ∀ z a b : × , 0 = 1 := by
intro z a b
@ -61,13 +60,14 @@ lemma derivSymm (f : ) (h : Differentiable f) :
dsimp [f'', f'] at A
apply A
theorem re_comp_holomorphic_is_harmonic (f : ) :
theorem holomorphic_is_harmonic (f : ) :
Differentiable f → Harmonic f := by
intro h
constructor
· -- Complex.reCLM ∘ f is two times real continuously differentiable
· -- f is two times real continuously differentiable
exact ContDiff.restrict_scalars (Differentiable.contDiff h)
· -- Laplace of f is zero