@@ -113,12 +113,11 @@ noncomputable theory
113
113
114
114
universes u v w u' v' w'
115
115
116
- open set filter
116
+ open set filter function
117
117
open_locale manifold filter topological_space
118
118
119
119
localized " notation `∞` := (⊤ : with_top ℕ)" in manifold
120
120
121
- section model_with_corners
122
121
/-! ### Models with corners. -/
123
122
124
123
/-- 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
213
212
@[simp, mfld_simps] protected lemma left_inv (x : H) : I.symm (I x) = x :=
214
213
by { refine I.left_inv' _, simp }
215
214
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
217
219
218
220
@[simp, mfld_simps] lemma symm_comp_self : I.symm ∘ I = id :=
219
221
I.left_inverse.comp_eq_id
542
544
543
545
end cont_diff_groupoid
544
546
545
- end model_with_corners
546
-
547
547
section smooth_manifold_with_corners
548
548
549
549
/-! ### Smooth manifolds with corners -/
@@ -603,9 +603,9 @@ def maximal_atlas := (cont_diff_groupoid ∞ I).maximal_atlas M
603
603
604
604
variable {M}
605
605
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 _
609
609
610
610
lemma chart_mem_maximal_atlas [smooth_manifold_with_corners I M] (x : M) :
611
611
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 }
708
708
lemma mem_ext_chart_source : x ∈ (ext_chart_at I x).source :=
709
709
by simp only [ext_chart_at_source, mem_chart_source]
710
710
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
+
711
716
lemma ext_chart_at_to_inv :
712
717
(ext_chart_at I x).symm ((ext_chart_at I x) x) = x :=
713
718
(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) :
873
878
𝓝[(ext_chart_at I x).symm ⁻¹' s ∩ range I] ((ext_chart_at I x) x) :=
874
879
ext_chart_preimage_mem_nhds_within' I x (mem_ext_chart_source I x) ht
875
880
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
+
876
888
/-- Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point
877
889
is a neighborhood of the preimage. -/
878
890
lemma ext_chart_preimage_mem_nhds (ht : t ∈ 𝓝 x) :
@@ -889,11 +901,41 @@ lemma ext_chart_preimage_inter_eq :
889
901
= ((ext_chart_at I x).symm ⁻¹' s ∩ range I) ∩ ((ext_chart_at I x).symm ⁻¹' t) :=
890
902
by mfld_set_tac
891
903
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
893
932
894
933
/-- In the case of the manifold structure on a vector space, the extended charts are just the
895
934
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 :=
899
936
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
0 commit comments