This commit is contained in:
Stefan Kebekus 2024-05-31 09:21:35 +02:00
parent c8e1aacb15
commit a7b0790675
2 changed files with 50 additions and 24 deletions

View File

@ -53,35 +53,45 @@ theorem laplace_add {f₁ f₂ : → F} (h₁ : ContDiff 2 f₁) (h₂ :
exact h₂.differentiable one_le_two
theorem laplace_add_ContDiffOn {f₁ f₂ : → F} {s : Set } (hs : IsOpen s) (h₁ : ContDiffOn 2 f₁ s) (h₂ : ContDiffOn 2 f₂ s): ∀ x ∈ s, Complex.laplace (f₁ + f₂) x = (Complex.laplace f₁) x + (Complex.laplace f₂) x := by
theorem laplace_add_ContDiffOn
{f₁ f₂ : → F}
{s : Set }
(hs : IsOpen s)
(h₁ : ContDiffOn 2 f₁ s)
(h₂ : ContDiffOn 2 f₂ s) :
∀ x ∈ s, Complex.laplace (f₁ + f₂) x = (Complex.laplace f₁) x + (Complex.laplace f₂) x := by
unfold Complex.laplace
simp
intro x hx
have : partialDeriv 1 (f₁ + f₂) =ᶠ[nhds x] (partialDeriv 1 f₁) + (partialDeriv 1 f₂) := by
sorry
rw [partialDeriv_eventuallyEq this]
have t₁ : DifferentiableAt (partialDeriv 1 f₁) x := by
sorry
have t₂ : DifferentiableAt (partialDeriv 1 f₂) x := by
sorry
rw [partialDeriv_add₂_differentiableAt t₁ t₂]
have : partialDeriv Complex.I (f₁ + f₂) =ᶠ[nhds x] (partialDeriv Complex.I f₁) + (partialDeriv Complex.I f₂) := by
sorry
rw [partialDeriv_eventuallyEq this]
have t₃ : DifferentiableAt (partialDeriv Complex.I f₁) x := by
sorry
have t₄ : DifferentiableAt (partialDeriv Complex.I f₂) x := by
sorry
rw [partialDeriv_add₂_differentiableAt t₃ t₄]
rw [partialDeriv_add₂]
rw [partialDeriv_add₂]
rw [partialDeriv_add₂]
rw [partialDeriv_add₂]
exact
add_add_add_comm (partialDeriv 1 (partialDeriv 1 f₁))
(partialDeriv 1 (partialDeriv 1 f₂))
(partialDeriv Complex.I (partialDeriv Complex.I f₁))
(partialDeriv Complex.I (partialDeriv Complex.I f₂))
exact (partialDeriv_contDiff h₁ Complex.I).differentiable le_rfl
exact (partialDeriv_contDiff h₂ Complex.I).differentiable le_rfl
exact h₁.differentiable one_le_two
exact h₂.differentiable one_le_two
exact (partialDeriv_contDiff h₁ 1).differentiable le_rfl
exact (partialDeriv_contDiff h₂ 1).differentiable le_rfl
exact h₁.differentiable one_le_two
exact h₂.differentiable one_le_two
-- I am super confused at this point because the tactic 'ring' does not work.
-- I do not understand why.
rw [add_assoc]
rw [add_assoc]
rw [add_right_inj (partialDeriv 1 (partialDeriv 1 f₁) x)]
rw [add_comm]
rw [add_assoc]
rw [add_right_inj (partialDeriv Complex.I (partialDeriv Complex.I f₁) x)]
rw [add_comm]
theorem laplace_smul {f : → F} (h : ContDiff 2 f) : ∀ v : , Complex.laplace (v • f) = v • (Complex.laplace f) := by

View File

@ -54,6 +54,22 @@ theorem partialDeriv_add₂ {f₁ f₂ : E → F} {v : E} (h₁ : Differentiable
rw [fderiv_add (h₁ w) (h₂ w)]
theorem partialDeriv_add₂_differentiableAt
{f₁ f₂ : E → F}
{v : E}
{x : E}
(h₁ : DifferentiableAt 𝕜 f₁ x)
(h₂ : DifferentiableAt 𝕜 f₂ x)
:
partialDeriv 𝕜 v (f₁ + f₂) x = (partialDeriv 𝕜 v f₁) x + (partialDeriv 𝕜 v f₂) x := by
unfold partialDeriv
have : f₁ + f₂ = fun y ↦ f₁ y + f₂ y := by rfl
rw [this]
rw [fderiv_add h₁ h₂]
rfl
theorem partialDeriv_compContLin {f : E → F} {l : F →L[𝕜] G} {v : E} (h : Differentiable 𝕜 f) : partialDeriv 𝕜 v (l ∘ f) = l ∘ partialDeriv 𝕜 v f := by
unfold partialDeriv