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

Commit 4ece2ee

Browse files
committed
feat(geometry/manifold/smooth_manifold_with_corners): properties to prove smoothness of mfderiv (#15523)
* Define `achart` * Rename and reformulate `mem_maximal_atlas_of_mem_atlas -> subset_maximal_atlas` * Used to prove smoothness of `mfderiv` (that result still has to be generalized to `mfderiv_within`, so is not included yet) * From the sphere eversion project
1 parent 0271f81 commit 4ece2ee

File tree

4 files changed

+82
-18
lines changed

4 files changed

+82
-18
lines changed

src/geometry/manifold/charted_space.lean

Lines changed: 15 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -503,6 +503,16 @@ lemma chart_source_mem_nhds (x : M) : (chart_at H x).source ∈ 𝓝 x :=
503503
lemma chart_target_mem_nhds (x : M) : (chart_at H x).target ∈ 𝓝 (chart_at H x x) :=
504504
(chart_at H x).open_target.mem_nhds $ mem_chart_target H x
505505

506+
/-- `achart H x` is the chart at `x`, considered as an element of the atlas.
507+
Especially useful for working with `basic_smooth_vector_bundle_core` -/
508+
def achart (x : M) : atlas H M := ⟨chart_at H x, chart_mem_atlas H x⟩
509+
510+
lemma achart_def (x : M) : achart H x = ⟨chart_at H x, chart_mem_atlas H x⟩ := rfl
511+
@[simp, mfld_simps]
512+
lemma coe_achart (x : M) : (achart H x : local_homeomorph M H) = chart_at H x := rfl
513+
@[simp, mfld_simps]
514+
lemma achart_val (x : M) : (achart H x).1 = chart_at H x := rfl
515+
506516
open topological_space
507517

508518
lemma charted_space.second_countable_of_countable_cover [second_countable_topology H]
@@ -792,13 +802,13 @@ def structure_groupoid.maximal_atlas : set (local_homeomorph M H) :=
792802
variable {M}
793803

794804
/-- The elements of the atlas belong to the maximal atlas for any structure groupoid -/
795-
lemma structure_groupoid.mem_maximal_atlas_of_mem_atlas [has_groupoid M G]
796-
{e : local_homeomorph M H} (he : e ∈ atlas H M) : e ∈ G.maximal_atlas M :=
797-
λ e' he', ⟨G.compatible he he', G.compatible he' he⟩
805+
lemma structure_groupoid.subset_maximal_atlas [has_groupoid M G] :
806+
atlas H M G.maximal_atlas M :=
807+
λ e he e' he', ⟨G.compatible he he', G.compatible he' he⟩
798808

799809
lemma structure_groupoid.chart_mem_maximal_atlas [has_groupoid M G]
800810
(x : M) : chart_at H x ∈ G.maximal_atlas M :=
801-
G.mem_maximal_atlas_of_mem_atlas (chart_mem_atlas H x)
811+
G.subset_maximal_atlas (chart_mem_atlas H x)
802812

803813
variable {G}
804814

@@ -835,7 +845,7 @@ variable (G)
835845

836846
/-- In the model space, the identity is in any maximal atlas. -/
837847
lemma structure_groupoid.id_mem_maximal_atlas : local_homeomorph.refl H ∈ G.maximal_atlas H :=
838-
G.mem_maximal_atlas_of_mem_atlas (by simp)
848+
G.subset_maximal_atlas $ by simp
839849

840850
end maximal_atlas
841851

src/geometry/manifold/cont_mdiff.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -719,6 +719,18 @@ lemma cont_mdiff_at_iff_cont_mdiff_on_nhds {n : ℕ} :
719719
by simp [← cont_mdiff_within_at_univ, cont_mdiff_within_at_iff_cont_mdiff_on_nhds,
720720
nhds_within_univ]
721721

722+
/-- Note: This does not hold for `n = ∞`. `f` being `C^∞` at `x` means that for every `n`, `f` is
723+
`C^n` on some neighborhood of `x`, but this neighborhood can depend on `n`. -/
724+
lemma cont_mdiff_at_iff_cont_mdiff_at_nhds {n : ℕ} :
725+
cont_mdiff_at I I' n f x ↔ ∀ᶠ x' in 𝓝 x, cont_mdiff_at I I' n f x' :=
726+
begin
727+
refine ⟨_, λ h, h.self_of_nhds⟩,
728+
rw [cont_mdiff_at_iff_cont_mdiff_on_nhds],
729+
rintro ⟨u, hu, h⟩,
730+
refine (eventually_mem_nhds.mpr hu).mono (λ x' hx', _),
731+
exact (h x' $ mem_of_mem_nhds hx').cont_mdiff_at hx'
732+
end
733+
722734
omit Is I's
723735

724736
/-! ### Congruence lemmas -/

src/geometry/manifold/smooth_manifold_with_corners.lean

Lines changed: 54 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -113,12 +113,11 @@ noncomputable theory
113113

114114
universes u v w u' v' w'
115115

116-
open set filter
116+
open set filter function
117117
open_locale manifold filter topological_space
118118

119119
localized "notation `∞` := (⊤ : with_top ℕ)" in manifold
120120

121-
section model_with_corners
122121
/-! ### Models with corners. -/
123122

124123
/-- A structure containing informations on the way a space `H` embeds in a
@@ -213,7 +212,10 @@ protected lemma unique_diff : unique_diff_on 𝕜 (range I) := I.target_eq ▸ I
213212
@[simp, mfld_simps] protected lemma left_inv (x : H) : I.symm (I x) = x :=
214213
by { refine I.left_inv' _, simp }
215214

216-
protected lemma left_inverse : function.left_inverse I.symm I := I.left_inv
215+
protected lemma left_inverse : left_inverse I.symm I := I.left_inv
216+
217+
lemma injective : injective I :=
218+
I.left_inverse.injective
217219

218220
@[simp, mfld_simps] lemma symm_comp_self : I.symm ∘ I = id :=
219221
I.left_inverse.comp_eq_id
@@ -542,8 +544,6 @@ end
542544

543545
end cont_diff_groupoid
544546

545-
end model_with_corners
546-
547547
section smooth_manifold_with_corners
548548

549549
/-! ### Smooth manifolds with corners -/
@@ -603,9 +603,9 @@ def maximal_atlas := (cont_diff_groupoid ∞ I).maximal_atlas M
603603

604604
variable {M}
605605

606-
lemma mem_maximal_atlas_of_mem_atlas [smooth_manifold_with_corners I M]
607-
{e : local_homeomorph M H} (he : e ∈ atlas H M) : e ∈ maximal_atlas I M :=
608-
structure_groupoid.mem_maximal_atlas_of_mem_atlas _ he
606+
lemma subset_maximal_atlas [smooth_manifold_with_corners I M] :
607+
atlas H M maximal_atlas I M :=
608+
structure_groupoid.subset_maximal_atlas _
609609

610610
lemma chart_mem_maximal_atlas [smooth_manifold_with_corners I M] (x : M) :
611611
chart_at H x ∈ maximal_atlas I M :=
@@ -708,6 +708,11 @@ by { rw ext_chart_at_source, exact (chart_at H x).open_source }
708708
lemma mem_ext_chart_source : x ∈ (ext_chart_at I x).source :=
709709
by simp only [ext_chart_at_source, mem_chart_source]
710710

711+
lemma ext_chart_at_target (x : M) : (ext_chart_at I x).target =
712+
I.symm ⁻¹' (chart_at H x).target ∩ range I :=
713+
by simp_rw [ext_chart_at, local_equiv.trans_target, I.target_eq, I.to_local_equiv_coe_symm,
714+
inter_comm]
715+
711716
lemma ext_chart_at_to_inv :
712717
(ext_chart_at I x).symm ((ext_chart_at I x) x) = x :=
713718
(ext_chart_at I x).left_inv (mem_ext_chart_source I x)
@@ -873,6 +878,13 @@ lemma ext_chart_preimage_mem_nhds_within (ht : t ∈ 𝓝[s] x) :
873878
𝓝[(ext_chart_at I x).symm ⁻¹' s ∩ range I] ((ext_chart_at I x) x) :=
874879
ext_chart_preimage_mem_nhds_within' I x (mem_ext_chart_source I x) ht
875880

881+
lemma ext_chart_preimage_mem_nhds' {x' : M} (h : x' ∈ (ext_chart_at I x).source) (ht : t ∈ 𝓝 x') :
882+
(ext_chart_at I x).symm ⁻¹' t ∈ 𝓝 (ext_chart_at I x x') :=
883+
begin
884+
apply (ext_chart_continuous_at_symm' I x h).preimage_mem_nhds,
885+
rwa (ext_chart_at I x).left_inv h
886+
end
887+
876888
/-- Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point
877889
is a neighborhood of the preimage. -/
878890
lemma ext_chart_preimage_mem_nhds (ht : t ∈ 𝓝 x) :
@@ -889,11 +901,41 @@ lemma ext_chart_preimage_inter_eq :
889901
= ((ext_chart_at I x).symm ⁻¹' s ∩ range I) ∩ ((ext_chart_at I x).symm ⁻¹' t) :=
890902
by mfld_set_tac
891903

892-
end extended_charts
904+
/-! We use the name `ext_coord_change` for `(ext_chart_at I x').symm ≫ ext_chart_at I x`. -/
905+
906+
lemma ext_coord_change_source (x x' : M) :
907+
((ext_chart_at I x').symm ≫ ext_chart_at I x).source =
908+
I '' ((chart_at H x').symm ≫ₕ (chart_at H x)).source :=
909+
by { simp_rw [local_equiv.trans_source, I.image_eq, ext_chart_at_source, local_equiv.symm_source,
910+
ext_chart_at_target, inter_right_comm _ (range I)], refl }
911+
912+
lemma cont_diff_on_ext_coord_change [smooth_manifold_with_corners I M] (x x' : M) :
913+
cont_diff_on 𝕜 ⊤ (ext_chart_at I x ∘ (ext_chart_at I x').symm)
914+
((ext_chart_at I x').symm ≫ ext_chart_at I x).source :=
915+
by { rw [ext_coord_change_source, I.image_eq], exact (has_groupoid.compatible
916+
(cont_diff_groupoid ⊤ I) (chart_mem_atlas H x') (chart_mem_atlas H x)).1 }
917+
918+
lemma cont_diff_within_at_ext_coord_change [smooth_manifold_with_corners I M] (x x' : M) {y : E}
919+
(hy : y ∈ ((ext_chart_at I x').symm ≫ ext_chart_at I x).source) :
920+
cont_diff_within_at 𝕜 ⊤ (ext_chart_at I x ∘ (ext_chart_at I x').symm) (range I) y :=
921+
begin
922+
apply (cont_diff_on_ext_coord_change I x x' y hy).mono_of_mem,
923+
rw [ext_coord_change_source] at hy ⊢,
924+
obtain ⟨z, hz, rfl⟩ := hy,
925+
exact I.image_mem_nhds_within ((local_homeomorph.open_source _).mem_nhds hz)
926+
end
927+
928+
variable (𝕜)
929+
930+
lemma ext_chart_self_eq {x : H} : ⇑(ext_chart_at I x) = I := rfl
931+
lemma ext_chart_self_apply {x y : H} : ext_chart_at I x y = I y := rfl
893932

894933
/-- In the case of the manifold structure on a vector space, the extended charts are just the
895934
identity.-/
896-
lemma ext_chart_model_space_eq_id (𝕜 : Type*) [nontrivially_normed_field 𝕜]
897-
{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] (x : E) :
898-
ext_chart_at (model_with_corners_self 𝕜 E) x = local_equiv.refl E :=
935+
lemma ext_chart_model_space_eq_id (x : E) : ext_chart_at 𝓘(𝕜, E) x = local_equiv.refl E :=
899936
by simp only with mfld_simps
937+
938+
lemma ext_chart_model_space_apply {x y : E} : ext_chart_at 𝓘(𝕜, E) x y = y := rfl
939+
940+
941+
end extended_charts

src/geometry/manifold/tangent_bundle.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -153,7 +153,7 @@ end
153153
def to_topological_vector_bundle_core : topological_vector_bundle_core 𝕜 M F (atlas H M) :=
154154
{ base_set := λ i, i.1.source,
155155
is_open_base_set := λ i, i.1.open_source,
156-
index_at := λ x, ⟨chart_at H x, chart_mem_atlas H x⟩,
156+
index_at := achart H,
157157
mem_base_set_at := λ x, mem_chart_source H x,
158158
coord_change := λ i j x, Z.coord_change i j (i.1 x),
159159
coord_change_self := λ i x hx v, Z.coord_change_self i (i.1 x) (i.1.map_source hx) v,

0 commit comments

Comments
 (0)