diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index d0411a26962e4..d7aea69808316 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -2867,4 +2867,37 @@ lemma differentiable.restrict_scalars (h : differentiable π•œ' f) : differentiable π•œ f := Ξ»x, (h x).restrict_scalars π•œ +lemma has_fderiv_within_at_of_restrict_scalars + {g' : E β†’L[π•œ] F} (h : has_fderiv_within_at f g' s x) + (H : f'.restrict_scalars π•œ = g') : has_fderiv_within_at f f' s x := +by { rw ← H at h, exact h } + +lemma has_fderiv_at_of_restrict_scalars {g' : E β†’L[π•œ] F} (h : has_fderiv_at f g' x) + (H : f'.restrict_scalars π•œ = g') : has_fderiv_at f f' x := +by { rw ← H at h, exact h } + +lemma differentiable_at.fderiv_restrict_scalars (h : differentiable_at π•œ' f x) : + fderiv π•œ f x = (fderiv π•œ' f x).restrict_scalars π•œ := +(h.has_fderiv_at.restrict_scalars π•œ).fderiv + +lemma differentiable_within_at_iff_restrict_scalars + (hf : differentiable_within_at π•œ f s x) (hs : unique_diff_within_at π•œ s x) : + differentiable_within_at π•œ' f s x ↔ + βˆƒ (g' : E β†’L[π•œ'] F), g'.restrict_scalars π•œ = fderiv_within π•œ f s x := +begin + split, + { rintros ⟨g', hg'⟩, + exact ⟨g', hs.eq (hg'.restrict_scalars π•œ) hf.has_fderiv_within_at⟩, }, + { rintros ⟨f', hf'⟩, + exact ⟨f', has_fderiv_within_at_of_restrict_scalars π•œ hf.has_fderiv_within_at hf'⟩, }, +end + +lemma differentiable_at_iff_restrict_scalars (hf : differentiable_at π•œ f x) : + differentiable_at π•œ' f x ↔ βˆƒ (g' : E β†’L[π•œ'] F), g'.restrict_scalars π•œ = fderiv π•œ f x := +begin + rw [← differentiable_within_at_univ, ← fderiv_within_univ], + exact differentiable_within_at_iff_restrict_scalars π•œ + hf.differentiable_within_at unique_diff_within_at_univ, +end + end restrict_scalars