@@ -23,8 +23,8 @@ for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`.
23
23
partition of unity, smooth bump function, whitney theorem
24
24
-/
25
25
26
- universes uE uF uH uM
27
- variables
26
+ universes uι uE uH uM
27
+ variables {ι : Type uι}
28
28
{E : Type uE} [normed_group E] [normed_space ℝ E] [finite_dimensional ℝ E]
29
29
{H : Type uH} [topological_space H] {I : model_with_corners ℝ E H}
30
30
{M : Type uM} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M]
@@ -36,78 +36,81 @@ noncomputable theory
36
36
37
37
namespace smooth_bump_covering
38
38
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
+ -/
40
45
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)
43
47
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),
47
51
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) }
49
53
50
54
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) :=
52
56
rfl
53
57
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 :=
55
59
begin
56
60
intros x hx y hy h,
57
61
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₁
63
67
end
64
68
65
- lemma embedding_pi_tangent_injective :
69
+ lemma embedding_pi_tangent_injective (f : smooth_bump_covering ι I M) :
66
70
injective f.embedding_pi_tangent :=
67
71
injective_iff_inj_on_univ.2 f.embedding_pi_tangent_inj_on
68
72
69
73
lemma comp_embedding_pi_tangent_mfderiv (x : M) (hx : x ∈ s) :
70
74
((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 :=
75
79
begin
76
80
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,
79
83
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, _),
82
86
simp only [embedding_pi_tangent_coe, continuous_linear_map.coe_comp', (∘),
83
87
continuous_linear_map.coe_fst', continuous_linear_map.proj_apply],
84
88
rw [hy, pi.one_apply, one_smul]
85
89
end
86
90
87
91
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 = ⊥ :=
89
93
begin
90
94
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],
93
97
exact linear_map.ker_le_ker_comp _ _
94
98
end
95
99
96
100
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)
99
103
100
104
/-- Baby version of the Whitney weak embedding theorem: if `M` admits a finite covering by
101
105
supports of bump functions, then for some `n` it can be immersed into the `n`-dimensional
102
106
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) :
105
108
∃ (n : ℕ) (e : M → euclidean_space ℝ (fin n)), smooth I (𝓡 n) e ∧
106
109
injective e ∧ ∀ x : M, injective (mfderiv I (𝓡 n) e x) :=
107
110
begin
108
- set F := euclidean_space ℝ (fin $ finrank ℝ (f. ι → (E × ℝ))),
111
+ set F := euclidean_space ℝ (fin $ finrank ℝ (ι → (E × ℝ))),
109
112
letI : finite_dimensional ℝ (E × ℝ) := by apply_instance,
110
- set eEF : (f. ι → (E × ℝ)) ≃L[ℝ] F :=
113
+ set eEF : (ι → (E × ℝ)) ≃L[ℝ] F :=
111
114
continuous_linear_equiv.of_finrank_eq finrank_euclidean_space_fin.symm,
112
115
refine ⟨_, eEF ∘ f.embedding_pi_tangent,
113
116
eEF.to_diffeomorph.smooth.comp f.embedding_pi_tangent.smooth,
@@ -122,11 +125,13 @@ end smooth_bump_covering
122
125
/-- Baby version of the Whitney weak embedding theorem: if `M` admits a finite covering by
123
126
supports of bump functions, then for some `n` it can be embedded into the `n`-dimensional
124
127
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] :
126
129
∃ (n : ℕ) (e : M → euclidean_space ℝ (fin n)), smooth I (𝓡 n) e ∧
127
130
closed_embedding e ∧ ∀ x : M, injective (mfderiv I (𝓡 n) e x) :=
128
131
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⟩,
131
136
exact ⟨n, e, hsmooth, hsmooth.continuous.closed_embedding hinj, hinj_mfderiv⟩
132
137
end
0 commit comments