Fix errors

This commit is contained in:
Stefan Kebekus 2024-05-17 09:19:24 +02:00
parent aeda1e981d
commit bcb639a5be
2 changed files with 10 additions and 14 deletions

View File

@ -44,7 +44,7 @@ theorem CauchyRiemann₃ : (DifferentiableAt f z)
simp simp
theorem CauchyRiemann₄ {f : } : (Differentiable f) theorem CauchyRiemann₄ {F : Type*} [NormedAddCommGroup F] [NormedSpace F] {f : → F} : (Differentiable f)
→ partialDeriv Complex.I f = Complex.I • partialDeriv 1 f := by → partialDeriv Complex.I f = Complex.I • partialDeriv 1 f := by
intro h intro h
unfold partialDeriv unfold partialDeriv

View File

@ -26,20 +26,20 @@ def Harmonic (f : → F) : Prop :=
(ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0) (ContDiff 2 f) ∧ (∀ z, Complex.laplace f z = 0)
theorem harmonic_comp_CLM_is_harmonic {f : → F₁} {l : F₁ →L[] G} (h : Harmonic f) : theorem harmonic_comp_CLM_is_harmonic {f : → F₁} {l : F₁ →L[] G} (h : Harmonic f) :
Harmonic (l ∘ f) := by Harmonic (l ∘ f) := by
constructor constructor
· -- Continuous differentiability · -- Continuous differentiability
apply ContDiff.comp apply ContDiff.comp
exact ContinuousLinearMap.contDiff l exact ContinuousLinearMap.contDiff l
exact ContDiff.restrict_scalars (Differentiable.contDiff h) exact h.1
· rw [laplace_compContLin] · rw [laplace_compContLin]
simp simp
intro z intro z
rw [h.2 z] rw [h.2 z]
simp simp
exact ContDiff.restrict_scalars (Differentiable.contDiff h) exact ContDiff.restrict_scalars h.1
theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[] G₁} : theorem harmonic_iff_comp_CLE_is_harmonic {f : → F₁} {l : F₁ ≃L[] G₁} :
@ -78,21 +78,17 @@ theorem holomorphic_is_harmonic {f : → F₁} (h : Differentiable f) :
-- This lemma says that partial derivatives commute with complex scalar -- This lemma says that partial derivatives commute with complex scalar
-- multiplication. This is a consequence of partialDeriv_compContLin once we -- multiplication. This is a consequence of partialDeriv_compContLin once we
-- note that complex scalar multiplication is continuous -linear. -- note that complex scalar multiplication is continuous -linear.
have : ∀ v, ∀ s : , ∀ g : , Differentiable g → partialDeriv v (s • g) = s • (partialDeriv v g) := by have : ∀ v, ∀ s : , ∀ g : F₁, Differentiable g → partialDeriv v (s • g) = s • (partialDeriv v g) := by
intro v s g hg intro v s g hg
-- Present scalar multiplication as a continuous -linear map. This is -- Present scalar multiplication as a continuous -linear map. This is
-- horrible, there must be better ways to do that. -- horrible, there must be better ways to do that.
let sMuls : →L[] := let sMuls : F₁ →L[] F₁ :=
{ {
toFun := fun x ↦ s * x toFun := fun x ↦ s • x
map_add' := by map_add' := by exact fun x y => DistribSMul.smul_add s x y
intro x y map_smul' := by exact fun m x => (smul_comm ((RingHom.id ) m) s x).symm
ring cont := continuous_const_smul s
map_smul' := by
intro m x
simp
ring
} }
-- Bring the goal into a form that is recognized by -- Bring the goal into a form that is recognized by