working…
This commit is contained in:
@@ -219,6 +219,7 @@ theorem laplace_compCLMAt
|
||||
obtain ⟨u, hu₁, hu₂⟩ : ∃ u ∈ nhds x, ContDiffOn ℝ 2 f u := by
|
||||
apply ContDiffAt.contDiffOn h
|
||||
rfl
|
||||
simp
|
||||
obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁
|
||||
use v
|
||||
constructor
|
||||
|
@@ -281,7 +281,11 @@ lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) :
|
||||
unfold partialDeriv
|
||||
rw [fderiv_clm_apply]
|
||||
· simp
|
||||
· exact (contDiff_succ_iff_fderiv.1 hf).2.differentiable le_rfl z
|
||||
· apply Differentiable.differentiableAt
|
||||
rw [← one_add_one_eq_two] at hf
|
||||
rw [contDiff_succ_iff_fderiv] at hf
|
||||
apply hf.2.2.differentiable
|
||||
simp
|
||||
· simp
|
||||
|
||||
|
||||
@@ -298,7 +302,9 @@ lemma partialDeriv_fderivOn
|
||||
· simp
|
||||
· convert DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz)
|
||||
apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1)
|
||||
exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 hf).2
|
||||
rw [← one_add_one_eq_two] at hf
|
||||
rw [contDiffOn_succ_iff_fderiv_of_isOpen hs] at hf
|
||||
exact hf.2.2
|
||||
· simp
|
||||
|
||||
|
||||
@@ -407,7 +413,8 @@ theorem partialDeriv_comm
|
||||
let f'' := (fderiv ℝ f' z)
|
||||
have h₁ : HasFDerivAt f' f'' z := by
|
||||
apply DifferentiableAt.hasFDerivAt
|
||||
apply (contDiff_succ_iff_fderiv.1 h).right.differentiable (Preorder.le_refl 1)
|
||||
rw [← one_add_one_eq_two] at h
|
||||
apply (contDiff_succ_iff_fderiv.1 h).2.2.differentiable (Preorder.le_refl 1)
|
||||
|
||||
apply second_derivative_symmetric h₀ h₁ v₁ v₂
|
||||
|
||||
@@ -442,7 +449,8 @@ theorem partialDeriv_commOn
|
||||
apply DifferentiableAt.hasFDerivAt
|
||||
apply DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz)
|
||||
apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1)
|
||||
exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 h).2
|
||||
rw [← one_add_one_eq_two] at h
|
||||
exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 h).2.2
|
||||
|
||||
have h₀' : ∀ᶠ (y : E) in nhds z, HasFDerivAt f (f' y) y := by
|
||||
apply eventually_nhds_iff.mpr
|
||||
@@ -463,7 +471,9 @@ theorem partialDeriv_commAt
|
||||
(h : ContDiffAt ℝ 2 f z) :
|
||||
∀ v₁ v₂ : E, partialDeriv ℝ v₁ (partialDeriv ℝ v₂ f) z = partialDeriv ℝ v₂ (partialDeriv ℝ v₁ f) z := by
|
||||
|
||||
obtain ⟨u, hu₁, hu₂⟩ := h.contDiffOn le_rfl
|
||||
let A := h.contDiffOn le_rfl
|
||||
simp at A
|
||||
obtain ⟨u, hu₁, hu₂⟩ := A
|
||||
obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁
|
||||
|
||||
intro v₁ v₂
|
||||
|
Reference in New Issue
Block a user