File tree Expand file tree Collapse file tree 1 file changed +13
-2
lines changed
Mathlib/Geometry/Manifold/ContMDiff Expand file tree Collapse file tree 1 file changed +13
-2
lines changed Original file line number Diff line number Diff line change @@ -340,15 +340,26 @@ theorem smoothWithinAt_one [One M'] : SmoothWithinAt I I' (1 : M → M') s x :=
340
340
341
341
end id
342
342
343
+ /-- `f` is continuously differentiable if it is cont. differentiable at each `x ∈ tsupport f`. -/
343
344
theorem contMDiff_of_support {f : M → F} (hf : ∀ x ∈ tsupport f, ContMDiffAt I 𝓘(𝕜, F) n f x) :
344
345
ContMDiff I 𝓘(𝕜, F) n f := by
345
346
intro x
346
347
by_cases hx : x ∈ tsupport f
347
348
· exact hf x hx
348
- · refine' ContMDiffAt.congr_of_eventuallyEq _ (eventuallyEq_zero_nhds.2 hx)
349
- exact contMDiffAt_const
349
+ · exact ContMDiffAt.congr_of_eventuallyEq contMDiffAt_const (eventuallyEq_zero_nhds.2 hx)
350
350
#align cont_mdiff_of_support contMDiff_of_support
351
351
352
+ theorem contMDiffWithinAt_of_not_mem {f : M → F} {x : M} (hx : x ∉ tsupport f) (n : ℕ∞)
353
+ (s : Set M) : ContMDiffWithinAt I 𝓘(𝕜, F) n f s x :=
354
+ contMDiffWithinAt_const.congr_of_eventuallyEq
355
+ (eventually_nhdsWithin_of_eventually_nhds <| not_mem_tsupport_iff_eventuallyEq.mp hx)
356
+ (image_eq_zero_of_nmem_tsupport hx)
357
+
358
+ /-- `f` is continuously differentiable at each point outside of its `tsupport`. -/
359
+ theorem contMDiffAt_of_not_mem {f : M → F} {x : M} (hx : x ∉ tsupport f) (n : ℕ∞) :
360
+ ContMDiffAt I 𝓘(𝕜, F) n f x :=
361
+ contMDiffWithinAt_of_not_mem hx n univ
362
+
352
363
/-! ### The inclusion map from one open set to another is smooth -/
353
364
354
365
section Inclusion
You can’t perform that action at this time.
0 commit comments