Fix errors
This commit is contained in:
parent
aeda1e981d
commit
bcb639a5be
|
@ -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
|
||||||
|
|
|
@ -33,13 +33,13 @@ theorem harmonic_comp_CLM_is_harmonic {f : ℂ → F₁} {l : F₁ →L[ℝ] G}
|
||||||
· -- 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
|
||||||
|
|
Loading…
Reference in New Issue