nevanlinna/Nevanlinna/partialDeriv.lean

45 lines
1.3 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Data.Fin.Tuple.Basic
import Mathlib.Analysis.Complex.Basic
import Mathlib.Analysis.Complex.TaylorSeries
import Mathlib.Analysis.Calculus.LineDeriv.Basic
import Mathlib.Analysis.Calculus.ContDiff.Basic
import Mathlib.Analysis.Calculus.ContDiff.Defs
import Mathlib.Analysis.Calculus.FDeriv.Basic
import Mathlib.Analysis.Calculus.FDeriv.Symmetric
noncomputable def Real.partialDeriv : → () → () := by
intro v
intro f
exact fun w ↦ (fderiv f w) v
theorem partialDeriv_smul {f : } {a v : } (h : Differentiable f) : Real.partialDeriv v (a • f) = a • Real.partialDeriv v f := by
unfold Real.partialDeriv
have : a • f = fun y ↦ a • f y := by rfl
rw [this]
conv =>
left
intro w
rw [fderiv_const_smul (h w)]
theorem partialDeriv_contDiff {n : } {f : } (h : ContDiff (n + 1) f) : ∀ v : , ContDiff n (Real.partialDeriv v f) := by
unfold Real.partialDeriv
intro v
let A := (contDiff_succ_iff_fderiv.1 h).right
simp at A
have : (fun w => (fderiv f w) v) = (fun f => f v) ∘ (fun w => (fderiv f w)) := by
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