Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit f471b89

Browse files
committed
feat(topology,geometry/manifold): continuous and smooth partition of unity (#8281)
Fixes #6392
1 parent dc6adcc commit f471b89

File tree

7 files changed

+800
-111
lines changed

7 files changed

+800
-111
lines changed

src/geometry/manifold/bump_function.lean

Lines changed: 8 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -13,18 +13,18 @@ import geometry.manifold.instances.real
1313
In this file we define `smooth_bump_function I c` to be a bundled smooth "bump" function centered at
1414
`c`. It is a structure that consists of two real numbers `0 < r < R` with small enough `R`. We
1515
define a coercion to function for this type, and for `f : smooth_bump_function I c`, the function
16-
`⇑f` written in the extended chart at `f.c` has the following properties:
16+
`⇑f` written in the extended chart at `c` has the following properties:
1717
18-
* `f x = 1` in the closed euclidean ball of radius `f.r` centered at `f.c`;
19-
* `f x = 0` outside of the euclidean ball of radius `f.R` centered at `f.c`;
18+
* `f x = 1` in the closed euclidean ball of radius `f.r` centered at `c`;
19+
* `f x = 0` outside of the euclidean ball of radius `f.R` centered at `c`;
2020
* `0 ≤ f x ≤ 1` for all `x`.
2121
22-
The actual statements involve (pre)images under `ext_chart_at I f.c` and are given as lemmas in the
22+
The actual statements involve (pre)images under `ext_chart_at I f` and are given as lemmas in the
2323
`smooth_bump_function` namespace.
2424
2525
## Tags
2626
27-
manifold, smooth bump function, partition of unity, Whitney theorem
27+
manifold, smooth bump function
2828
-/
2929

3030
universes uE uF uH uM
@@ -54,7 +54,7 @@ In this section we define a structure for a bundled smooth bump function and pro
5454
5555
The structure contains data required to construct a function with these properties. The function is
5656
available as `⇑f` or `f x`. Formal statements of the properties listed above involve some
57-
(pre)images under `ext_chart_at I f.c` and are given as lemmas in the `msmooth_bump_function`
57+
(pre)images under `ext_chart_at I f.c` and are given as lemmas in the `smooth_bump_function`
5858
namespace. -/
5959
structure smooth_bump_function (c : M) extends times_cont_diff_bump (ext_chart_at I c c) :=
6060
(closed_ball_subset :
@@ -307,6 +307,8 @@ end
307307

308308
protected lemma smooth_at {x} : smooth_at I 𝓘(ℝ) f x := f.smooth.smooth_at
309309

310+
protected lemma continuous : continuous f := f.smooth.continuous
311+
310312
/-- If `f : smooth_bump_function I c` is a smooth bump function and `g : M → G` is a function smooth
311313
on the source of the chart at `c`, then `f • g` is smooth on the whole manifold. -/
312314
lemma smooth_smul {G} [normed_group G] [normed_space ℝ G]

src/geometry/manifold/partition_of_unity.lean

Lines changed: 277 additions & 65 deletions
Large diffs are not rendered by default.

src/geometry/manifold/times_cont_mdiff.lean

Lines changed: 33 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -923,26 +923,59 @@ begin
923923
exact times_cont_diff_within_at_const,
924924
end
925925

926+
@[to_additive]
927+
lemma times_cont_mdiff_one [has_one M'] : times_cont_mdiff I I' n (1 : M → M') :=
928+
by simp only [pi.one_def, times_cont_mdiff_const]
929+
926930
lemma smooth_const : smooth I I' (λ (x : M), c) := times_cont_mdiff_const
927931

932+
@[to_additive]
933+
lemma smooth_one [has_one M'] : smooth I I' (1 : M → M') :=
934+
by simp only [pi.one_def, smooth_const]
935+
928936
lemma times_cont_mdiff_on_const : times_cont_mdiff_on I I' n (λ (x : M), c) s :=
929937
times_cont_mdiff_const.times_cont_mdiff_on
930938

939+
@[to_additive]
940+
lemma times_cont_mdiff_on_one [has_one M'] : times_cont_mdiff_on I I' n (1 : M → M') s :=
941+
times_cont_mdiff_one.times_cont_mdiff_on
942+
931943
lemma smooth_on_const : smooth_on I I' (λ (x : M), c) s :=
932944
times_cont_mdiff_on_const
933945

946+
@[to_additive]
947+
lemma smooth_on_one [has_one M'] : smooth_on I I' (1 : M → M') s :=
948+
times_cont_mdiff_on_one
949+
934950
lemma times_cont_mdiff_at_const : times_cont_mdiff_at I I' n (λ (x : M), c) x :=
935951
times_cont_mdiff_const.times_cont_mdiff_at
936952

953+
@[to_additive]
954+
lemma times_cont_mdiff_at_one [has_one M'] : times_cont_mdiff_at I I' n (1 : M → M') x :=
955+
times_cont_mdiff_one.times_cont_mdiff_at
956+
937957
lemma smooth_at_const : smooth_at I I' (λ (x : M), c) x :=
938958
times_cont_mdiff_at_const
939959

960+
@[to_additive]
961+
lemma smooth_at_one [has_one M'] : smooth_at I I' (1 : M → M') x :=
962+
times_cont_mdiff_at_one
963+
940964
lemma times_cont_mdiff_within_at_const : times_cont_mdiff_within_at I I' n (λ (x : M), c) s x :=
941965
times_cont_mdiff_at_const.times_cont_mdiff_within_at
942966

967+
@[to_additive]
968+
lemma times_cont_mdiff_within_at_one [has_one M'] :
969+
times_cont_mdiff_within_at I I' n (1 : M → M') s x :=
970+
times_cont_mdiff_at_const.times_cont_mdiff_within_at
971+
943972
lemma smooth_within_at_const : smooth_within_at I I' (λ (x : M), c) s x :=
944973
times_cont_mdiff_within_at_const
945974

975+
@[to_additive]
976+
lemma smooth_within_at_one [has_one M'] : smooth_within_at I I' (1 : M → M') s x :=
977+
times_cont_mdiff_within_at_one
978+
946979
end id
947980

948981
lemma times_cont_mdiff_of_support {f : M → F}

src/geometry/manifold/times_cont_mdiff_map.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -56,6 +56,10 @@ attribute [to_additive_ignore_args 21] times_cont_mdiff_map
5656
times_cont_mdiff_map.has_coe_to_fun times_cont_mdiff_map.continuous_map.has_coe
5757
variables {f g : C^n⟮I, M; I', M'⟯}
5858

59+
@[simp] lemma coe_fn_mk (f : M → M') (hf : times_cont_mdiff I I' n f) :
60+
(mk f hf : M → M') = f :=
61+
rfl
62+
5963
protected lemma times_cont_mdiff (f : C^n⟮I, M; I', M'⟯) :
6064
times_cont_mdiff I I' n f := f.times_cont_mdiff_to_fun
6165

src/geometry/manifold/whitney_embedding.lean

Lines changed: 42 additions & 37 deletions
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,8 @@ for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`.
2323
partition of unity, smooth bump function, whitney theorem
2424
-/
2525

26-
universes uE uF uH uM
27-
variables
26+
universes uι uE uH uM
27+
variables {ι : Type uι}
2828
{E : Type uE} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E]
2929
{H : Type uH} [topological_space H] {I : model_with_corners ℝ E H}
3030
{M : Type uM} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M]
@@ -36,78 +36,81 @@ noncomputable theory
3636

3737
namespace smooth_bump_covering
3838

39-
variables {s : set M} {U : M → set M} (fs : smooth_bump_covering I s)
39+
/-!
40+
### Whitney embedding theorem
41+
42+
In this section we prove a version of the Whitney embedding theorem: for any compact real manifold
43+
`M`, for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`.
44+
-/
4045

41-
variables [t2_space M] [fintype fs.ι] (f : smooth_bump_covering I (univ : set M))
42-
[fintype f.ι]
46+
variables [t2_space M] [fintype ι] {s : set M} (f : smooth_bump_covering ι I M s)
4347

44-
/-- Smooth embedding of `M` into `(E × ℝ) ^ f.ι`. -/
45-
def embedding_pi_tangent : C^∞⟮I, M; 𝓘(ℝ, fs.ι → (E × ℝ)), fs.ι → (E × ℝ)⟯ :=
46-
{ to_fun := λ x i, (fs i x • ext_chart_at I (fs.c i) x, fs i x),
48+
/-- Smooth embedding of `M` into `(E × ℝ) ^ ι`. -/
49+
def embedding_pi_tangent : C^∞⟮I, M; 𝓘(ℝ, ι → (E × ℝ)), ι → (E × ℝ)⟯ :=
50+
{ to_fun := λ x i, (f i x • ext_chart_at I (f.c i) x, f i x),
4751
times_cont_mdiff_to_fun := times_cont_mdiff_pi_space.2 $ λ i,
48-
((fs i).smooth_smul times_cont_mdiff_on_ext_chart_at).prod_mk_space ((fs i).smooth) }
52+
((f i).smooth_smul times_cont_mdiff_on_ext_chart_at).prod_mk_space ((f i).smooth) }
4953

5054
local attribute [simp] lemma embedding_pi_tangent_coe :
51-
fs.embedding_pi_tangent = λ x i, (fs i x • ext_chart_at I (fs.c i) x, fs i x) :=
55+
f.embedding_pi_tangent = λ x i, (f i x • ext_chart_at I (f.c i) x, f i x) :=
5256
rfl
5357

54-
lemma embedding_pi_tangent_inj_on : inj_on fs.embedding_pi_tangent s :=
58+
lemma embedding_pi_tangent_inj_on : inj_on f.embedding_pi_tangent s :=
5559
begin
5660
intros x hx y hy h,
5761
simp only [embedding_pi_tangent_coe, funext_iff] at h,
58-
obtain ⟨h₁, h₂⟩ := prod.mk.inj_iff.1 (h (fs.ind x hx)),
59-
rw [fs.apply_ind x hx] at h₂,
60-
rw [← h₂, fs.apply_ind x hx, one_smul, one_smul] at h₁,
61-
have := fs.mem_ext_chart_at_source_of_eq_one h₂.symm,
62-
exact (ext_chart_at I (fs.c _)).inj_on (fs.mem_ext_chart_at_ind_source x hx) this h₁
62+
obtain ⟨h₁, h₂⟩ := prod.mk.inj_iff.1 (h (f.ind x hx)),
63+
rw [f.apply_ind x hx] at h₂,
64+
rw [← h₂, f.apply_ind x hx, one_smul, one_smul] at h₁,
65+
have := f.mem_ext_chart_at_source_of_eq_one h₂.symm,
66+
exact (ext_chart_at I (f.c _)).inj_on (f.mem_ext_chart_at_ind_source x hx) this h₁
6367
end
6468

65-
lemma embedding_pi_tangent_injective :
69+
lemma embedding_pi_tangent_injective (f : smooth_bump_covering ι I M) :
6670
injective f.embedding_pi_tangent :=
6771
injective_iff_inj_on_univ.2 f.embedding_pi_tangent_inj_on
6872

6973
lemma comp_embedding_pi_tangent_mfderiv (x : M) (hx : x ∈ s) :
7074
((continuous_linear_map.fst ℝ E ℝ).comp
71-
(@continuous_linear_map.proj ℝ _ fs.ι (λ _, E × ℝ) _ _
72-
(λ _, infer_instance) (fs.ind x hx))).comp
73-
(mfderiv I 𝓘(ℝ, fs.ι → (E × ℝ)) fs.embedding_pi_tangent x) =
74-
mfderiv I I (chart_at H (fs.c (fs.ind x hx))) x :=
75+
(@continuous_linear_map.proj ℝ _ ι (λ _, E × ℝ) _ _
76+
(λ _, infer_instance) (f.ind x hx))).comp
77+
(mfderiv I 𝓘(ℝ, ι → (E × ℝ)) f.embedding_pi_tangent x) =
78+
mfderiv I I (chart_at H (f.c (f.ind x hx))) x :=
7579
begin
7680
set L := ((continuous_linear_map.fst ℝ E ℝ).comp
77-
(@continuous_linear_map.proj ℝ _ fs.ι (λ _, E × ℝ) _ _ (λ _, infer_instance) (fs.ind x hx))),
78-
have := (L.has_mfderiv_at.comp x (fs.embedding_pi_tangent.mdifferentiable_at.has_mfderiv_at)),
81+
(@continuous_linear_map.proj ℝ _ ι (λ _, E × ℝ) _ _ (λ _, infer_instance) (f.ind x hx))),
82+
have := L.has_mfderiv_at.comp x f.embedding_pi_tangent.mdifferentiable_at.has_mfderiv_at,
7983
convert has_mfderiv_at_unique this _,
80-
refine (has_mfderiv_at_ext_chart_at I (fs.mem_chart_at_ind_source x hx)).congr_of_eventually_eq _,
81-
refine (fs.eventually_eq_one x hx).mono (λ y hy, _),
84+
refine (has_mfderiv_at_ext_chart_at I (f.mem_chart_at_ind_source x hx)).congr_of_eventually_eq _,
85+
refine (f.eventually_eq_one x hx).mono (λ y hy, _),
8286
simp only [embedding_pi_tangent_coe, continuous_linear_map.coe_comp', (∘),
8387
continuous_linear_map.coe_fst', continuous_linear_map.proj_apply],
8488
rw [hy, pi.one_apply, one_smul]
8589
end
8690

8791
lemma embedding_pi_tangent_ker_mfderiv (x : M) (hx : x ∈ s) :
88-
(mfderiv I 𝓘(ℝ, fs.ι → (E × ℝ)) fs.embedding_pi_tangent x).ker = ⊥ :=
92+
(mfderiv I 𝓘(ℝ, ι → (E × ℝ)) f.embedding_pi_tangent x).ker = ⊥ :=
8993
begin
9094
apply bot_unique,
91-
rw [← (mdifferentiable_chart I (fs.c (fs.ind x hx))).ker_mfderiv_eq_bot
92-
(fs.mem_chart_at_ind_source x hx), ← comp_embedding_pi_tangent_mfderiv],
95+
rw [← (mdifferentiable_chart I (f.c (f.ind x hx))).ker_mfderiv_eq_bot
96+
(f.mem_chart_at_ind_source x hx), ← comp_embedding_pi_tangent_mfderiv],
9397
exact linear_map.ker_le_ker_comp _ _
9498
end
9599

96100
lemma embedding_pi_tangent_injective_mfderiv (x : M) (hx : x ∈ s) :
97-
injective (mfderiv I 𝓘(ℝ, fs.ι → (E × ℝ)) fs.embedding_pi_tangent x) :=
98-
linear_map.ker_eq_bot.1 (fs.embedding_pi_tangent_ker_mfderiv x hx)
101+
injective (mfderiv I 𝓘(ℝ, ι → (E × ℝ)) f.embedding_pi_tangent x) :=
102+
linear_map.ker_eq_bot.1 (f.embedding_pi_tangent_ker_mfderiv x hx)
99103

100104
/-- Baby version of the Whitney weak embedding theorem: if `M` admits a finite covering by
101105
supports of bump functions, then for some `n` it can be immersed into the `n`-dimensional
102106
Euclidean space. -/
103-
lemma exists_immersion_finrank (f : smooth_bump_covering I (univ : set M))
104-
[fintype f.ι] :
107+
lemma exists_immersion_euclidean (f : smooth_bump_covering ι I M) :
105108
∃ (n : ℕ) (e : M → euclidean_space ℝ (fin n)), smooth I (𝓡 n) e ∧
106109
injective e ∧ ∀ x : M, injective (mfderiv I (𝓡 n) e x) :=
107110
begin
108-
set F := euclidean_space ℝ (fin $ finrank ℝ (f.ι → (E × ℝ))),
111+
set F := euclidean_space ℝ (fin $ finrank ℝ (ι → (E × ℝ))),
109112
letI : finite_dimensional ℝ (E × ℝ) := by apply_instance,
110-
set eEF : (f.ι → (E × ℝ)) ≃L[ℝ] F :=
113+
set eEF : (ι → (E × ℝ)) ≃L[ℝ] F :=
111114
continuous_linear_equiv.of_finrank_eq finrank_euclidean_space_fin.symm,
112115
refine ⟨_, eEF ∘ f.embedding_pi_tangent,
113116
eEF.to_diffeomorph.smooth.comp f.embedding_pi_tangent.smooth,
@@ -122,11 +125,13 @@ end smooth_bump_covering
122125
/-- Baby version of the Whitney weak embedding theorem: if `M` admits a finite covering by
123126
supports of bump functions, then for some `n` it can be embedded into the `n`-dimensional
124127
Euclidean space. -/
125-
lemma exists_embedding_finrank_of_compact [t2_space M] [compact_space M] :
128+
lemma exists_embedding_euclidean_of_compact [t2_space M] [compact_space M] :
126129
∃ (n : ℕ) (e : M → euclidean_space ℝ (fin n)), smooth I (𝓡 n) e ∧
127130
closed_embedding e ∧ ∀ x : M, injective (mfderiv I (𝓡 n) e x) :=
128131
begin
129-
rcases (smooth_bump_covering.choice I M).exists_immersion_finrank
130-
with ⟨n, e, hsmooth, hinj, hinj_mfderiv⟩,
132+
rcases smooth_bump_covering.exists_is_subordinate I is_closed_univ (λ (x : M) _, univ_mem_sets)
133+
with ⟨ι, f, -⟩,
134+
haveI := f.fintype,
135+
rcases f.exists_immersion_euclidean with ⟨n, e, hsmooth, hinj, hinj_mfderiv⟩,
131136
exact ⟨n, e, hsmooth, hsmooth.continuous.closed_embedding hinj, hinj_mfderiv⟩
132137
end

src/topology/paracompact.lean

Lines changed: 1 addition & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -38,9 +38,7 @@ We also prove the following facts.
3838
3939
## TODO
4040
41-
* Define partition of unity.
42-
43-
* Prove (some of) [Michael's theorems](https://ncatlab.org/nlab/show/Michael%27s+theorem).
41+
Prove (some of) [Michael's theorems](https://ncatlab.org/nlab/show/Michael%27s+theorem).
4442
4543
## Tags
4644

0 commit comments

Comments
 (0)