Skip to content

Commit 0bcd6be

Browse files
committed
feat: two lemmas about cut-off functions (#9873)
From sphere-eversion; I'm just upstreaming this. The version in sphere-eversion uses an unbundled design; we provide a bundled version (for now) to match the remaining file. Co-authored-by: grunweg <grunweg@posteo.de>
1 parent f1919fd commit 0bcd6be

File tree

1 file changed

+30
-1
lines changed

1 file changed

+30
-1
lines changed

Mathlib/Geometry/Manifold/PartitionOfUnity.lean

Lines changed: 30 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -483,7 +483,7 @@ See also `exists_msmooth_zero_iff_one_iff_of_isClosed`, which ensures additional
483483
`f` is equal to `0` exactly on `s` and to `1` exactly on `t`. -/
484484
theorem exists_smooth_zero_one_of_isClosed [T2Space M] [SigmaCompactSpace M] {s t : Set M}
485485
(hs : IsClosed s) (ht : IsClosed t) (hd : Disjoint s t) :
486-
∃ f : C^∞⟮I, M; 𝓘(ℝ), ℝ⟯, EqOn f 0 s ∧ EqOn f 1 t ∧ ∀ x, f x ∈ Icc (0 : ℝ) 1 := by
486+
∃ f : C^∞⟮I, M; 𝓘(ℝ), ℝ⟯, EqOn f 0 s ∧ EqOn f 1 t ∧ ∀ x, f x ∈ Icc 0 1 := by
487487
have : ∀ x ∈ t, sᶜ ∈ 𝓝 x := fun x hx => hs.isOpen_compl.mem_nhds (disjoint_right.1 hd hx)
488488
rcases SmoothBumpCovering.exists_isSubordinate I ht this with ⟨ι, f, hf⟩
489489
set g := f.toSmoothPartitionOfUnity
@@ -495,6 +495,35 @@ theorem exists_smooth_zero_one_of_isClosed [T2Space M] [SigmaCompactSpace M] {s
495495
exact nmem_support.1 (subset_compl_comm.1 (hf.support_subset i) hx)
496496
#align exists_smooth_zero_one_of_closed exists_smooth_zero_one_of_isClosed
497497

498+
/-- Given two disjoint closed sets `s, t` in a Hausdorff normal σ-compact finite dimensional
499+
manifold `M`, there exists a smooth function `f : M → [0,1]` that vanishes in a neighbourhood of `s`
500+
and is equal to `1` in a neighbourhood of `t`. -/
501+
theorem exists_smooth_zero_one_nhds_of_isClosed [T2Space M] [NormalSpace M] [SigmaCompactSpace M]
502+
{s t : Set M} (hs : IsClosed s) (ht : IsClosed t) (hd : Disjoint s t) :
503+
∃ f : C^∞⟮I, M; 𝓘(ℝ), ℝ⟯, (∀ᶠ x in 𝓝ˢ s, f x = 0) ∧ (∀ᶠ x in 𝓝ˢ t, f x = 1) ∧
504+
∀ x, f x ∈ Icc 0 1 := by
505+
obtain ⟨u, u_op, hsu, hut⟩ := normal_exists_closure_subset hs ht.isOpen_compl
506+
(subset_compl_iff_disjoint_left.mpr hd.symm)
507+
obtain ⟨v, v_op, htv, hvu⟩ := normal_exists_closure_subset ht isClosed_closure.isOpen_compl
508+
(subset_compl_comm.mp hut)
509+
obtain ⟨f, hfu, hfv, hf⟩ := exists_smooth_zero_one_of_isClosed I isClosed_closure isClosed_closure
510+
(subset_compl_iff_disjoint_left.mp hvu)
511+
refine ⟨f, ?_, ?_, hf⟩
512+
· exact eventually_of_mem (mem_of_superset (u_op.mem_nhdsSet.mpr hsu) subset_closure) hfu
513+
· exact eventually_of_mem (mem_of_superset (v_op.mem_nhdsSet.mpr htv) subset_closure) hfv
514+
515+
/-- Given two sets `s, t` in a Hausdorff normal σ-compact finite-dimensional manifold `M`
516+
with `s` open and `s ⊆ interior t`, there is a smooth function `f : M → [0,1]` which is equal to `s`
517+
in a neighbourhood of `s` and has support contained in `t`. -/
518+
theorem exists_smooth_one_nhds_of_subset_interior [T2Space M] [NormalSpace M] [SigmaCompactSpace M]
519+
{s t : Set M} (hs : IsClosed s) (hd : s ⊆ interior t) :
520+
∃ f : C^∞⟮I, M; 𝓘(ℝ), ℝ⟯, (∀ᶠ x in 𝓝ˢ s, f x = 1) ∧ (∀ x ∉ t, f x = 0) ∧
521+
∀ x, f x ∈ Icc 0 1 := by
522+
rcases exists_smooth_zero_one_nhds_of_isClosed I isOpen_interior.isClosed_compl hs
523+
(by rwa [← subset_compl_iff_disjoint_left, compl_compl]) with ⟨f, h0, h1, hf⟩
524+
refine ⟨f, h1, fun x hx ↦ ?_, hf⟩
525+
exact h0.self_of_nhdsSet _ fun hx' ↦ hx <| interior_subset hx'
526+
498527
namespace SmoothPartitionOfUnity
499528

500529
/-- A `SmoothPartitionOfUnity` that consists of a single function, uniformly equal to one,

0 commit comments

Comments
 (0)