Update partialDeriv.lean
This commit is contained in:
parent
3bdc7eaffb
commit
f480ae2a0f
|
@ -91,31 +91,26 @@ theorem partialDeriv_contDiff {n : ℕ} {f : E → F} (h : ContDiff 𝕜 (n + 1)
|
||||||
|
|
||||||
|
|
||||||
theorem partialDeriv_contDiffAt {n : ℕ} {f : E → F} {x : E} (h : ContDiffAt 𝕜 (n + 1) f x) : ∀ v : E, ContDiffAt 𝕜 n (partialDeriv 𝕜 v f) x := by
|
theorem partialDeriv_contDiffAt {n : ℕ} {f : E → F} {x : E} (h : ContDiffAt 𝕜 (n + 1) f x) : ∀ v : E, ContDiffAt 𝕜 n (partialDeriv 𝕜 v f) x := by
|
||||||
|
|
||||||
unfold partialDeriv
|
unfold partialDeriv
|
||||||
intro v
|
intro v
|
||||||
|
|
||||||
let A' := (contDiffAt_succ_iff_hasFDerivAt.1 h)
|
let eval_at_v : (E →L[𝕜] F) →L[𝕜] F :=
|
||||||
obtain ⟨f', ⟨u, hu₁, hu₂⟩ , hf'⟩ := A'
|
{
|
||||||
|
toFun := fun l ↦ l v
|
||||||
|
map_add' := by simp
|
||||||
|
map_smul' := by simp
|
||||||
|
}
|
||||||
|
|
||||||
have : (fun w => (fderiv 𝕜 f w) v) = (fun f => f v) ∘ (fun w => (fderiv 𝕜 f w)) := by
|
have : (fun w => (fderiv 𝕜 f w) v) = eval_at_v ∘ (fun w => (fderiv 𝕜 f w)) := by
|
||||||
rfl
|
rfl
|
||||||
rw [this]
|
rw [this]
|
||||||
apply ContDiffAt.comp
|
|
||||||
apply fderiv_clm_apply
|
|
||||||
|
|
||||||
let A := (contDiffAt_succ_iff_fderiv.1 h).right
|
apply ContDiffAt.continuousLinearMap_comp
|
||||||
simp at A
|
-- ContDiffAt 𝕜 (↑n) (fun w => fderiv 𝕜 f w) x
|
||||||
|
apply ContDiffAt.fderiv_right h
|
||||||
have : (fun w => (fderiv 𝕜 f w) v) = (fun f => f v) ∘ (fun w => (fderiv 𝕜 f w)) := by
|
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
rw [this]
|
|
||||||
refine ContDiff.comp ?hg A
|
|
||||||
refine ContDiff.of_succ ?hg.h
|
|
||||||
refine ContDiff.clm_apply ?hg.h.hf ?hg.h.hg
|
|
||||||
exact contDiff_id
|
|
||||||
exact contDiff_const
|
|
||||||
|
|
||||||
|
|
||||||
lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) :
|
lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) :
|
||||||
fderiv 𝕜 (fderiv 𝕜 f) z b a = partialDeriv 𝕜 b (partialDeriv 𝕜 a f) z := by
|
fderiv 𝕜 (fderiv 𝕜 f) z b a = partialDeriv 𝕜 b (partialDeriv 𝕜 a f) z := by
|
||||||
|
|
Loading…
Reference in New Issue