import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.FDeriv.Add variable {š•œ : Type*} [NontriviallyNormedField š•œ] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace š•œ E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace š•œ F] variable {G : Type*} [NormedAddCommGroup G] [NormedSpace š•œ G] variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace š•œ G'] variable {f fā‚€ f₁ g : E → F} variable {f' fā‚€' f₁' g' : E →L[š•œ] F} variable (e : E →L[š•œ] F) variable {x : E} variable {s t : Set E} variable {L L₁ Lā‚‚ : Filter E} variable {R : Type*} [Semiring R] [Module R F] [SMulCommClass š•œ R F] [ContinuousConstSMul R F] -- import Mathlib.Analysis.Calculus.FDeriv.Add @[fun_prop] theorem Differentiable.const_smul' (h : Differentiable š•œ f) (c : R) : Differentiable š•œ (c • f) := by have : c • f = fun x ↦ c • f x := rfl rw [this] exact Differentiable.const_smul h c -- Mathlib.Analysis.Calculus.ContDiff.Basic theorem ContDiff.const_smul' {f : E → F} (c : R) (hf : ContDiff š•œ n f) : ContDiff š•œ n (c • f) := by have : c • f = fun x ↦ c • f x := rfl rw [this] exact ContDiff.const_smul c hf