| author | wenzelm | 
| Thu, 09 Jan 2020 15:45:31 +0100 | |
| changeset 71360 | fcf5ee85743d | 
| parent 71184 | d62fdaafdafc | 
| child 71633 | 07bec530f02e | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Homeomorphism.thy | 
| 63130 | 2 | Author: LC Paulson (ported from HOL Light) | 
| 3 | *) | |
| 4 | ||
| 69722 
b5163b2132c5
minor tagging updates in 13 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 5 | section \<open>Homeomorphism Theorems\<close> | 
| 63130 | 6 | |
| 7 | theory Homeomorphism | |
| 69620 | 8 | imports Homotopy | 
| 63130 | 9 | begin | 
| 10 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 11 | lemma homeomorphic_spheres': | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 12 | fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 13 |   assumes "0 < \<delta>" and dimeq: "DIM('a) = DIM('b)"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 14 | shows "(sphere a \<delta>) homeomorphic (sphere b \<delta>)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 15 | proof - | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 16 | obtain f :: "'a\<Rightarrow>'b" and g where "linear f" "linear g" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 17 | and fg: "\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 18 | by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq]) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 19 | then have "continuous_on UNIV f" "continuous_on UNIV g" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 20 | using linear_continuous_on linear_linear by blast+ | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 21 | then show ?thesis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 22 | unfolding homeomorphic_minimal | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 23 | apply(rule_tac x="\<lambda>x. b + f(x - a)" in exI) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 24 | apply(rule_tac x="\<lambda>x. a + g(x - b)" in exI) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 25 | using assms | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 26 | apply (force intro: continuous_intros | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 27 | continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 28 | done | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 29 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 30 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 31 | lemma homeomorphic_spheres_gen: | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 32 | fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 33 |   assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 34 | shows "(sphere a r homeomorphic sphere b s)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 35 | apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres']) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 36 | using assms apply auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 37 | done | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64773diff
changeset | 38 | |
| 69683 | 39 | subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close> | 
| 63130 | 40 | |
| 69739 | 41 | proposition | 
| 63130 | 42 | fixes S :: "'a::euclidean_space set" | 
| 43 | assumes "compact S" and 0: "0 \<in> rel_interior S" | |
| 44 | and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment 0 x \<subseteq> rel_interior S" | |
| 45 | shows starlike_compact_projective1_0: | |
| 46 | "S - rel_interior S homeomorphic sphere 0 1 \<inter> affine hull S" | |
| 47 | (is "?SMINUS homeomorphic ?SPHER") | |
| 48 | and starlike_compact_projective2_0: | |
| 49 | "S homeomorphic cball 0 1 \<inter> affine hull S" | |
| 50 | (is "S homeomorphic ?CBALL") | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 51 | proof - | 
| 63130 | 52 | have starI: "(u *\<^sub>R x) \<in> rel_interior S" if "x \<in> S" "0 \<le> u" "u < 1" for x u | 
| 53 | proof (cases "x=0 \<or> u=0") | |
| 54 | case True with 0 show ?thesis by force | |
| 55 | next | |
| 56 | case False with that show ?thesis | |
| 57 | by (auto simp: in_segment intro: star [THEN subsetD]) | |
| 58 | qed | |
| 59 | have "0 \<in> S" using assms rel_interior_subset by auto | |
| 60 | define proj where "proj \<equiv> \<lambda>x::'a. x /\<^sub>R norm x" | |
| 61 | have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y | |
| 62 | using that by (force simp: proj_def) | |
| 63 | then have iff_eq: "\<And>x y. (proj x = proj y \<and> norm x = norm y) \<longleftrightarrow> x = y" | |
| 64 | by blast | |
| 65 | have projI: "x \<in> affine hull S \<Longrightarrow> proj x \<in> affine hull S" for x | |
| 66 | by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_mul proj_def) | |
| 67 | have nproj1 [simp]: "x \<noteq> 0 \<Longrightarrow> norm(proj x) = 1" for x | |
| 68 | by (simp add: proj_def) | |
| 69 | have proj0_iff [simp]: "proj x = 0 \<longleftrightarrow> x = 0" for x | |
| 70 | by (simp add: proj_def) | |
| 71 |   have cont_proj: "continuous_on (UNIV - {0}) proj"
 | |
| 72 | unfolding proj_def by (rule continuous_intros | force)+ | |
| 73 | have proj_spherI: "\<And>x. \<lbrakk>x \<in> affine hull S; x \<noteq> 0\<rbrakk> \<Longrightarrow> proj x \<in> ?SPHER" | |
| 74 | by (simp add: projI) | |
| 75 | have "bounded S" "closed S" | |
| 76 | using \<open>compact S\<close> compact_eq_bounded_closed by blast+ | |
| 77 | have inj_on_proj: "inj_on proj (S - rel_interior S)" | |
| 78 | proof | |
| 79 | fix x y | |
| 80 | assume x: "x \<in> S - rel_interior S" | |
| 81 | and y: "y \<in> S - rel_interior S" and eq: "proj x = proj y" | |
| 82 | then have xynot: "x \<noteq> 0" "y \<noteq> 0" "x \<in> S" "y \<in> S" "x \<notin> rel_interior S" "y \<notin> rel_interior S" | |
| 83 | using 0 by auto | |
| 84 | consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith | |
| 85 | then show "x = y" | |
| 86 | proof cases | |
| 87 | assume "norm x = norm y" | |
| 88 | with iff_eq eq show "x = y" by blast | |
| 89 | next | |
| 90 | assume *: "norm x < norm y" | |
| 91 | have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))" | |
| 92 | by force | |
| 93 | then have "proj ((norm x / norm y) *\<^sub>R y) = proj x" | |
| 94 | by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) | |
| 95 | then have [simp]: "(norm x / norm y) *\<^sub>R y = x" | |
| 96 | by (rule eqI) (simp add: \<open>y \<noteq> 0\<close>) | |
| 97 | have no: "0 \<le> norm x / norm y" "norm x / norm y < 1" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 98 | using * by (auto simp: field_split_simps) | 
| 63130 | 99 | then show "x = y" | 
| 100 | using starI [OF \<open>y \<in> S\<close> no] xynot by auto | |
| 101 | next | |
| 102 | assume *: "norm x > norm y" | |
| 103 | have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))" | |
| 104 | by force | |
| 105 | then have "proj ((norm y / norm x) *\<^sub>R x) = proj y" | |
| 106 | by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) | |
| 107 | then have [simp]: "(norm y / norm x) *\<^sub>R x = y" | |
| 108 | by (rule eqI) (simp add: \<open>x \<noteq> 0\<close>) | |
| 109 | have no: "0 \<le> norm y / norm x" "norm y / norm x < 1" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 110 | using * by (auto simp: field_split_simps) | 
| 63130 | 111 | then show "x = y" | 
| 112 | using starI [OF \<open>x \<in> S\<close> no] xynot by auto | |
| 113 | qed | |
| 114 | qed | |
| 115 | have "\<exists>surf. homeomorphism (S - rel_interior S) ?SPHER proj surf" | |
| 116 | proof (rule homeomorphism_compact) | |
| 117 | show "compact (S - rel_interior S)" | |
| 118 | using \<open>compact S\<close> compact_rel_boundary by blast | |
| 119 | show "continuous_on (S - rel_interior S) proj" | |
| 120 | using 0 by (blast intro: continuous_on_subset [OF cont_proj]) | |
| 121 | show "proj ` (S - rel_interior S) = ?SPHER" | |
| 122 | proof | |
| 123 | show "proj ` (S - rel_interior S) \<subseteq> ?SPHER" | |
| 124 | using 0 by (force simp: hull_inc projI intro: nproj1) | |
| 125 | show "?SPHER \<subseteq> proj ` (S - rel_interior S)" | |
| 126 | proof (clarsimp simp: proj_def) | |
| 127 | fix x | |
| 128 | assume "x \<in> affine hull S" and nox: "norm x = 1" | |
| 129 | then have "x \<noteq> 0" by auto | |
| 130 | obtain d where "0 < d" and dx: "(d *\<^sub>R x) \<in> rel_frontier S" | |
| 131 | and ri: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (e *\<^sub>R x) \<in> rel_interior S" | |
| 132 | using ray_to_rel_frontier [OF \<open>bounded S\<close> 0] \<open>x \<in> affine hull S\<close> \<open>x \<noteq> 0\<close> by auto | |
| 133 | show "x \<in> (\<lambda>x. x /\<^sub>R norm x) ` (S - rel_interior S)" | |
| 134 | apply (rule_tac x="d *\<^sub>R x" in image_eqI) | |
| 135 | using \<open>0 < d\<close> | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 136 | using dx \<open>closed S\<close> apply (auto simp: rel_frontier_def field_split_simps nox) | 
| 63130 | 137 | done | 
| 138 | qed | |
| 139 | qed | |
| 140 | qed (rule inj_on_proj) | |
| 141 | then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf" | |
| 142 | by blast | |
| 143 | then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf" | |
| 144 | by (auto simp: homeomorphism_def) | |
| 145 | have surf_nz: "\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0" | |
| 146 | by (metis "0" DiffE homeomorphism_def imageI surf) | |
| 147 | have cont_nosp: "continuous_on (?SPHER) (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))" | |
| 148 | apply (rule continuous_intros)+ | |
| 149 | apply (rule continuous_on_subset [OF cont_proj], force) | |
| 150 | apply (rule continuous_on_subset [OF cont_surf]) | |
| 151 | apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI) | |
| 152 | done | |
| 153 | have surfpS: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<in> S" | |
| 154 | by (metis (full_types) DiffE \<open>0 \<in> S\<close> homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf) | |
| 155 | have *: "\<exists>y. norm y = 1 \<and> y \<in> affine hull S \<and> x = surf (proj y)" | |
| 156 | if "x \<in> S" "x \<notin> rel_interior S" for x | |
| 157 | proof - | |
| 158 | have "proj x \<in> ?SPHER" | |
| 159 | by (metis (full_types) "0" hull_inc proj_spherI that) | |
| 160 | moreover have "surf (proj x) = x" | |
| 161 | by (metis Diff_iff homeomorphism_def surf that) | |
| 162 | ultimately show ?thesis | |
| 163 | by (metis \<open>\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0\<close> hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1)) | |
| 164 | qed | |
| 165 | have surfp_notin: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<notin> rel_interior S" | |
| 166 | by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf) | |
| 167 | have no_sp_im: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S" | |
| 168 | by (auto simp: surfpS image_def Bex_def surfp_notin *) | |
| 169 | have inj_spher: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?SPHER" | |
| 170 | proof | |
| 171 | fix x y | |
| 172 | assume xy: "x \<in> ?SPHER" "y \<in> ?SPHER" | |
| 173 | and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" | |
| 174 | then have "norm x = 1" "norm y = 1" "x \<in> affine hull S" "y \<in> affine hull S" | |
| 175 | using 0 by auto | |
| 176 | with eq show "x = y" | |
| 177 | by (simp add: proj_def) (metis surf xy homeomorphism_def) | |
| 178 | qed | |
| 179 | have co01: "compact ?SPHER" | |
| 71172 | 180 | by (simp add: compact_Int_closed) | 
| 63130 | 181 | show "?SMINUS homeomorphic ?SPHER" | 
| 182 | apply (subst homeomorphic_sym) | |
| 183 | apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher]) | |
| 184 | done | |
| 185 | have proj_scaleR: "\<And>a x. 0 < a \<Longrightarrow> proj (a *\<^sub>R x) = proj x" | |
| 186 | by (simp add: proj_def) | |
| 187 |   have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)"
 | |
| 188 | apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force) | |
| 189 | apply (rule continuous_on_subset [OF cont_surf]) | |
| 190 | using homeomorphism_image1 proj_spherI surf by fastforce | |
| 191 | obtain B where "B>0" and B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B" | |
| 192 | by (metis compact_imp_bounded \<open>compact S\<close> bounded_pos_less less_eq_real_def) | |
| 193 | have cont_nosp: "continuous (at x within ?CBALL) (\<lambda>x. norm x *\<^sub>R surf (proj x))" | |
| 194 | if "norm x \<le> 1" "x \<in> affine hull S" for x | |
| 195 | proof (cases "x=0") | |
| 196 | case True | |
| 197 | show ?thesis using True | |
| 198 | apply (simp add: continuous_within) | |
| 199 | apply (rule lim_null_scaleR_bounded [where B=B]) | |
| 200 | apply (simp_all add: tendsto_norm_zero eventually_at) | |
| 201 | apply (rule_tac x=B in exI) | |
| 202 | using B surfpS proj_def projI apply (auto simp: \<open>B > 0\<close>) | |
| 203 | done | |
| 204 | next | |
| 205 | case False | |
| 206 |     then have "\<forall>\<^sub>F x in at x. (x \<in> affine hull S - {0}) = (x \<in> affine hull S)"
 | |
| 207 | apply (simp add: eventually_at) | |
| 208 | apply (rule_tac x="norm x" in exI) | |
| 209 | apply (auto simp: False) | |
| 210 | done | |
| 211 | with cont_sp0 have *: "continuous (at x within affine hull S) (\<lambda>x. surf (proj x))" | |
| 212 | apply (simp add: continuous_on_eq_continuous_within) | |
| 213 | apply (drule_tac x=x in bspec, force simp: False that) | |
| 214 | apply (simp add: continuous_within Lim_transform_within_set) | |
| 215 | done | |
| 216 | show ?thesis | |
| 217 | apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2]) | |
| 218 | apply (rule continuous_intros *)+ | |
| 219 | done | |
| 220 | qed | |
| 221 | have cont_nosp2: "continuous_on ?CBALL (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))" | |
| 222 | by (simp add: continuous_on_eq_continuous_within cont_nosp) | |
| 223 | have "norm y *\<^sub>R surf (proj y) \<in> S" if "y \<in> cball 0 1" and yaff: "y \<in> affine hull S" for y | |
| 224 | proof (cases "y=0") | |
| 225 | case True then show ?thesis | |
| 226 | by (simp add: \<open>0 \<in> S\<close>) | |
| 227 | next | |
| 228 | case False | |
| 229 | then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))" | |
| 230 | by (simp add: proj_def) | |
| 231 | have "norm y \<le> 1" using that by simp | |
| 232 | have "surf (proj (y /\<^sub>R norm y)) \<in> S" | |
| 233 | apply (rule surfpS) | |
| 234 | using proj_def projI yaff | |
| 235 | by (auto simp: False) | |
| 236 | then have "surf (proj y) \<in> S" | |
| 237 | by (simp add: False proj_def) | |
| 238 | then show "norm y *\<^sub>R surf (proj y) \<in> S" | |
| 239 | by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one | |
| 240 | starI subset_eq \<open>norm y \<le> 1\<close>) | |
| 241 | qed | |
| 242 | moreover have "x \<in> (\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \<in> S" for x | |
| 243 | proof (cases "x=0") | |
| 244 | case True with that hull_inc show ?thesis by fastforce | |
| 245 | next | |
| 246 | case False | |
| 247 | then have psp: "proj (surf (proj x)) = proj x" | |
| 248 | by (metis homeomorphism_def hull_inc proj_spherI surf that) | |
| 249 | have nxx: "norm x *\<^sub>R proj x = x" | |
| 250 | by (simp add: False local.proj_def) | |
| 251 | have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \<in> affine hull S" | |
| 252 | by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_clauses(4) that) | |
| 253 | have sproj_nz: "surf (proj x) \<noteq> 0" | |
| 254 | by (metis False proj0_iff psp) | |
| 255 | then have "proj x = proj (proj x)" | |
| 256 | by (metis False nxx proj_scaleR zero_less_norm_iff) | |
| 257 | moreover have scaleproj: "\<And>a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a" | |
| 258 | by (simp add: divide_inverse local.proj_def) | |
| 259 | ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \<notin> rel_interior S" | |
| 260 | by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that) | |
| 261 | then have "(norm (surf (proj x)) / norm x) \<ge> 1" | |
| 262 | using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff) | |
| 263 | then have nole: "norm x \<le> norm (surf (proj x))" | |
| 264 | by (simp add: le_divide_eq_1) | |
| 265 | show ?thesis | |
| 266 | apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI) | |
| 267 | apply (metis (no_types, hide_lams) mult.commute scaleproj abs_inverse abs_norm_cancel divide_inverse norm_scaleR nxx positive_imp_inverse_positive proj_scaleR psp sproj_nz zero_less_norm_iff) | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 268 | apply (auto simp: field_split_simps nole affineI) | 
| 63130 | 269 | done | 
| 270 | qed | |
| 271 | ultimately have im_cball: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S" | |
| 272 | by blast | |
| 273 | have inj_cball: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?CBALL" | |
| 274 | proof | |
| 275 | fix x y | |
| 276 | assume "x \<in> ?CBALL" "y \<in> ?CBALL" | |
| 277 | and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" | |
| 278 | then have x: "x \<in> affine hull S" and y: "y \<in> affine hull S" | |
| 279 | using 0 by auto | |
| 280 | show "x = y" | |
| 281 | proof (cases "x=0 \<or> y=0") | |
| 282 | case True then show "x = y" using eq proj_spherI surf_nz x y by force | |
| 283 | next | |
| 284 | case False | |
| 285 | with x y have speq: "surf (proj x) = surf (proj y)" | |
| 286 | by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff) | |
| 287 | then have "norm x = norm y" | |
| 288 | by (metis \<open>x \<in> affine hull S\<close> \<open>y \<in> affine hull S\<close> eq proj_spherI real_vector.scale_cancel_right surf_nz) | |
| 289 | moreover have "proj x = proj y" | |
| 290 | by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y) | |
| 291 | ultimately show "x = y" | |
| 292 | using eq eqI by blast | |
| 293 | qed | |
| 294 | qed | |
| 295 | have co01: "compact ?CBALL" | |
| 71172 | 296 | by (simp add: compact_Int_closed) | 
| 63130 | 297 | show "S homeomorphic ?CBALL" | 
| 298 | apply (subst homeomorphic_sym) | |
| 299 | apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball]) | |
| 300 | done | |
| 301 | qed | |
| 302 | ||
| 69739 | 303 | corollary | 
| 63130 | 304 | fixes S :: "'a::euclidean_space set" | 
| 305 | assumes "compact S" and a: "a \<in> rel_interior S" | |
| 306 | and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" | |
| 307 | shows starlike_compact_projective1: | |
| 308 | "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S" | |
| 309 | and starlike_compact_projective2: | |
| 310 | "S homeomorphic cball a 1 \<inter> affine hull S" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 311 | proof - | 
| 67399 | 312 | have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation) | 
| 313 | have 2: "0 \<in> rel_interior ((+) (-a) ` S)" | |
| 69661 | 314 | using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp) | 
| 67399 | 315 | have 3: "open_segment 0 x \<subseteq> rel_interior ((+) (-a) ` S)" if "x \<in> ((+) (-a) ` S)" for x | 
| 63130 | 316 | proof - | 
| 317 | have "x+a \<in> S" using that by auto | |
| 318 | then have "open_segment a (x+a) \<subseteq> rel_interior S" by (metis star) | |
| 69661 | 319 | then show ?thesis using open_segment_translation [of a 0 x] | 
| 320 | using rel_interior_translation [of "- a" S] by (fastforce simp add: ac_simps image_iff cong: image_cong_simp) | |
| 63130 | 321 | qed | 
| 67399 | 322 | have "S - rel_interior S homeomorphic ((+) (-a) ` S) - rel_interior ((+) (-a) ` S)" | 
| 63130 | 323 | by (metis rel_interior_translation translation_diff homeomorphic_translation) | 
| 67399 | 324 | also have "... homeomorphic sphere 0 1 \<inter> affine hull ((+) (-a) ` S)" | 
| 63130 | 325 | by (rule starlike_compact_projective1_0 [OF 1 2 3]) | 
| 67399 | 326 | also have "... = (+) (-a) ` (sphere a 1 \<inter> affine hull S)" | 
| 63130 | 327 | by (metis affine_hull_translation left_minus sphere_translation translation_Int) | 
| 328 | also have "... homeomorphic sphere a 1 \<inter> affine hull S" | |
| 329 | using homeomorphic_translation homeomorphic_sym by blast | |
| 330 | finally show "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S" . | |
| 331 | ||
| 67399 | 332 | have "S homeomorphic ((+) (-a) ` S)" | 
| 63130 | 333 | by (metis homeomorphic_translation) | 
| 67399 | 334 | also have "... homeomorphic cball 0 1 \<inter> affine hull ((+) (-a) ` S)" | 
| 63130 | 335 | by (rule starlike_compact_projective2_0 [OF 1 2 3]) | 
| 67399 | 336 | also have "... = (+) (-a) ` (cball a 1 \<inter> affine hull S)" | 
| 63130 | 337 | by (metis affine_hull_translation left_minus cball_translation translation_Int) | 
| 338 | also have "... homeomorphic cball a 1 \<inter> affine hull S" | |
| 339 | using homeomorphic_translation homeomorphic_sym by blast | |
| 340 | finally show "S homeomorphic cball a 1 \<inter> affine hull S" . | |
| 341 | qed | |
| 342 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 343 | corollary starlike_compact_projective_special: | 
| 63130 | 344 | assumes "compact S" | 
| 345 | and cb01: "cball (0::'a::euclidean_space) 1 \<subseteq> S" | |
| 346 | and scale: "\<And>x u. \<lbrakk>x \<in> S; 0 \<le> u; u < 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x \<in> S - frontier S" | |
| 347 | shows "S homeomorphic (cball (0::'a::euclidean_space) 1)" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 348 | proof - | 
| 63130 | 349 | have "ball 0 1 \<subseteq> interior S" | 
| 350 | using cb01 interior_cball interior_mono by blast | |
| 351 | then have 0: "0 \<in> rel_interior S" | |
| 352 | by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le) | |
| 353 | have [simp]: "affine hull S = UNIV" | |
| 354 | using \<open>ball 0 1 \<subseteq> interior S\<close> by (auto intro!: affine_hull_nonempty_interior) | |
| 355 | have star: "open_segment 0 x \<subseteq> rel_interior S" if "x \<in> S" for x | |
| 63627 | 356 | proof | 
| 63130 | 357 | fix p assume "p \<in> open_segment 0 x" | 
| 358 | then obtain u where "x \<noteq> 0" and u: "0 \<le> u" "u < 1" and p: "u *\<^sub>R x = p" | |
| 63627 | 359 | by (auto simp: in_segment) | 
| 63130 | 360 | then show "p \<in> rel_interior S" | 
| 361 | using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce | |
| 362 | qed | |
| 363 | show ?thesis | |
| 364 | using starlike_compact_projective2_0 [OF \<open>compact S\<close> 0 star] by simp | |
| 365 | qed | |
| 366 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 367 | lemma homeomorphic_convex_lemma: | 
| 63130 | 368 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 369 | assumes "convex S" "compact S" "convex T" "compact T" | |
| 370 | and affeq: "aff_dim S = aff_dim T" | |
| 371 | shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \<and> | |
| 372 | S homeomorphic T" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 373 | proof (cases "rel_interior S = {} \<or> rel_interior T = {}")
 | 
| 63130 | 374 | case True | 
| 375 | then show ?thesis | |
| 376 | by (metis Diff_empty affeq \<open>convex S\<close> \<open>convex T\<close> aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty) | |
| 377 | next | |
| 378 | case False | |
| 379 | then obtain a b where a: "a \<in> rel_interior S" and b: "b \<in> rel_interior T" by auto | |
| 380 | have starS: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" | |
| 381 | using rel_interior_closure_convex_segment | |
| 382 | a \<open>convex S\<close> closure_subset subsetCE by blast | |
| 383 | have starT: "\<And>x. x \<in> T \<Longrightarrow> open_segment b x \<subseteq> rel_interior T" | |
| 384 | using rel_interior_closure_convex_segment | |
| 385 | b \<open>convex T\<close> closure_subset subsetCE by blast | |
| 67399 | 386 | let ?aS = "(+) (-a) ` S" and ?bT = "(+) (-b) ` T" | 
| 63130 | 387 | have 0: "0 \<in> affine hull ?aS" "0 \<in> affine hull ?bT" | 
| 388 | by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+ | |
| 389 | have subs: "subspace (span ?aS)" "subspace (span ?bT)" | |
| 390 | by (rule subspace_span)+ | |
| 391 | moreover | |
| 67399 | 392 | have "dim (span ((+) (- a) ` S)) = dim (span ((+) (- b) ` T))" | 
| 63130 | 393 | by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int) | 
| 394 | ultimately obtain f g where "linear f" "linear g" | |
| 395 | and fim: "f ` span ?aS = span ?bT" | |
| 396 | and gim: "g ` span ?bT = span ?aS" | |
| 397 | and fno: "\<And>x. x \<in> span ?aS \<Longrightarrow> norm(f x) = norm x" | |
| 398 | and gno: "\<And>x. x \<in> span ?bT \<Longrightarrow> norm(g x) = norm x" | |
| 399 | and gf: "\<And>x. x \<in> span ?aS \<Longrightarrow> g(f x) = x" | |
| 400 | and fg: "\<And>x. x \<in> span ?bT \<Longrightarrow> f(g x) = x" | |
| 401 | by (rule isometries_subspaces) blast | |
| 402 | have [simp]: "continuous_on A f" for A | |
| 403 | using \<open>linear f\<close> linear_conv_bounded_linear linear_continuous_on by blast | |
| 404 | have [simp]: "continuous_on B g" for B | |
| 405 | using \<open>linear g\<close> linear_conv_bounded_linear linear_continuous_on by blast | |
| 406 | have eqspanS: "affine hull ?aS = span ?aS" | |
| 407 | by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) | |
| 408 | have eqspanT: "affine hull ?bT = span ?bT" | |
| 409 | by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) | |
| 410 | have "S homeomorphic cball a 1 \<inter> affine hull S" | |
| 411 | by (rule starlike_compact_projective2 [OF \<open>compact S\<close> a starS]) | |
| 67399 | 412 | also have "... homeomorphic (+) (-a) ` (cball a 1 \<inter> affine hull S)" | 
| 63130 | 413 | by (metis homeomorphic_translation) | 
| 67399 | 414 | also have "... = cball 0 1 \<inter> (+) (-a) ` (affine hull S)" | 
| 63130 | 415 | by (auto simp: dist_norm) | 
| 416 | also have "... = cball 0 1 \<inter> span ?aS" | |
| 417 | using eqspanS affine_hull_translation by blast | |
| 418 | also have "... homeomorphic cball 0 1 \<inter> span ?bT" | |
| 419 | proof (rule homeomorphicI [where f=f and g=g]) | |
| 420 | show fim1: "f ` (cball 0 1 \<inter> span ?aS) = cball 0 1 \<inter> span ?bT" | |
| 421 | apply (rule subset_antisym) | |
| 422 | using fim fno apply (force simp:, clarify) | |
| 423 | by (metis IntI fg gim gno image_eqI mem_cball_0) | |
| 424 | show "g ` (cball 0 1 \<inter> span ?bT) = cball 0 1 \<inter> span ?aS" | |
| 425 | apply (rule subset_antisym) | |
| 426 | using gim gno apply (force simp:, clarify) | |
| 427 | by (metis IntI fim1 gf image_eqI) | |
| 428 | qed (auto simp: fg gf) | |
| 67399 | 429 | also have "... = cball 0 1 \<inter> (+) (-b) ` (affine hull T)" | 
| 63130 | 430 | using eqspanT affine_hull_translation by blast | 
| 67399 | 431 | also have "... = (+) (-b) ` (cball b 1 \<inter> affine hull T)" | 
| 63130 | 432 | by (auto simp: dist_norm) | 
| 433 | also have "... homeomorphic (cball b 1 \<inter> affine hull T)" | |
| 434 | by (metis homeomorphic_translation homeomorphic_sym) | |
| 435 | also have "... homeomorphic T" | |
| 436 | by (metis starlike_compact_projective2 [OF \<open>compact T\<close> b starT] homeomorphic_sym) | |
| 437 | finally have 1: "S homeomorphic T" . | |
| 438 | ||
| 439 | have "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S" | |
| 440 | by (rule starlike_compact_projective1 [OF \<open>compact S\<close> a starS]) | |
| 67399 | 441 | also have "... homeomorphic (+) (-a) ` (sphere a 1 \<inter> affine hull S)" | 
| 63130 | 442 | by (metis homeomorphic_translation) | 
| 67399 | 443 | also have "... = sphere 0 1 \<inter> (+) (-a) ` (affine hull S)" | 
| 63130 | 444 | by (auto simp: dist_norm) | 
| 445 | also have "... = sphere 0 1 \<inter> span ?aS" | |
| 446 | using eqspanS affine_hull_translation by blast | |
| 447 | also have "... homeomorphic sphere 0 1 \<inter> span ?bT" | |
| 448 | proof (rule homeomorphicI [where f=f and g=g]) | |
| 449 | show fim1: "f ` (sphere 0 1 \<inter> span ?aS) = sphere 0 1 \<inter> span ?bT" | |
| 450 | apply (rule subset_antisym) | |
| 451 | using fim fno apply (force simp:, clarify) | |
| 452 | by (metis IntI fg gim gno image_eqI mem_sphere_0) | |
| 453 | show "g ` (sphere 0 1 \<inter> span ?bT) = sphere 0 1 \<inter> span ?aS" | |
| 454 | apply (rule subset_antisym) | |
| 455 | using gim gno apply (force simp:, clarify) | |
| 456 | by (metis IntI fim1 gf image_eqI) | |
| 457 | qed (auto simp: fg gf) | |
| 67399 | 458 | also have "... = sphere 0 1 \<inter> (+) (-b) ` (affine hull T)" | 
| 63130 | 459 | using eqspanT affine_hull_translation by blast | 
| 67399 | 460 | also have "... = (+) (-b) ` (sphere b 1 \<inter> affine hull T)" | 
| 63130 | 461 | by (auto simp: dist_norm) | 
| 462 | also have "... homeomorphic (sphere b 1 \<inter> affine hull T)" | |
| 463 | by (metis homeomorphic_translation homeomorphic_sym) | |
| 464 | also have "... homeomorphic T - rel_interior T" | |
| 465 | by (metis starlike_compact_projective1 [OF \<open>compact T\<close> b starT] homeomorphic_sym) | |
| 466 | finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" . | |
| 467 | show ?thesis | |
| 468 | using 1 2 by blast | |
| 469 | qed | |
| 470 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 471 | lemma homeomorphic_convex_compact_sets: | 
| 63130 | 472 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 473 | assumes "convex S" "compact S" "convex T" "compact T" | |
| 474 | and affeq: "aff_dim S = aff_dim T" | |
| 475 | shows "S homeomorphic T" | |
| 476 | using homeomorphic_convex_lemma [OF assms] assms | |
| 477 | by (auto simp: rel_frontier_def) | |
| 478 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 479 | lemma homeomorphic_rel_frontiers_convex_bounded_sets: | 
| 63130 | 480 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 481 | assumes "convex S" "bounded S" "convex T" "bounded T" | |
| 482 | and affeq: "aff_dim S = aff_dim T" | |
| 483 | shows "rel_frontier S homeomorphic rel_frontier T" | |
| 484 | using assms homeomorphic_convex_lemma [of "closure S" "closure T"] | |
| 485 | by (simp add: rel_frontier_def convex_rel_interior_closure) | |
| 486 | ||
| 487 | ||
| 69683 | 488 | subsection\<open>Homeomorphisms between punctured spheres and affine sets\<close> | 
| 63130 | 489 | text\<open>Including the famous stereoscopic projection of the 3-D sphere to the complex plane\<close> | 
| 490 | ||
| 491 | text\<open>The special case with centre 0 and radius 1\<close> | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 492 | lemma homeomorphic_punctured_affine_sphere_affine_01: | 
| 63130 | 493 | assumes "b \<in> sphere 0 1" "affine T" "0 \<in> T" "b \<in> T" "affine p" | 
| 494 | and affT: "aff_dim T = aff_dim p + 1" | |
| 495 |     shows "(sphere 0 1 \<inter> T) - {b} homeomorphic p"
 | |
| 496 | proof - | |
| 497 | have [simp]: "norm b = 1" "b\<bullet>b = 1" | |
| 498 | using assms by (auto simp: norm_eq_1) | |
| 499 |   have [simp]: "T \<inter> {v. b\<bullet>v = 0} \<noteq> {}"
 | |
| 500 | using \<open>0 \<in> T\<close> by auto | |
| 501 |   have [simp]: "\<not> T \<subseteq> {v. b\<bullet>v = 0}"
 | |
| 502 | using \<open>norm b = 1\<close> \<open>b \<in> T\<close> by auto | |
| 503 | define f where "f \<equiv> \<lambda>x. 2 *\<^sub>R b + (2 / (1 - b\<bullet>x)) *\<^sub>R (x - b)" | |
| 504 | define g where "g \<equiv> \<lambda>y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)" | |
| 505 | have [simp]: "\<And>x. \<lbrakk>x \<in> T; b\<bullet>x = 0\<rbrakk> \<Longrightarrow> f (g x) = x" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 506 | unfolding f_def g_def by (simp add: algebra_simps field_split_simps add_nonneg_eq_0_iff) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 507 | have no: "(norm (f x))\<^sup>2 = 4 * (1 + b \<bullet> x) / (1 - b \<bullet> x)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 508 | if "norm x = 1" and "b \<bullet> x \<noteq> 1" for x | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 509 | using that | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 510 | apply (simp flip: dot_square_norm add: norm_eq_1 nonzero_eq_divide_eq) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 511 | apply (simp add: f_def vector_add_divide_simps inner_simps) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 512 | apply (use sum_sqs_eq [of 1 "b \<bullet> x"] | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 513 | in \<open>auto simp add: field_split_simps inner_commute\<close>) | 
| 63130 | 514 | done | 
| 515 | have [simp]: "\<And>u::real. 8 + u * (u * 8) = u * 16 \<longleftrightarrow> u=1" | |
| 516 | by algebra | |
| 517 | have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> g (f x) = x" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 518 | unfolding g_def no by (auto simp: f_def field_split_simps) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 519 | have [simp]: "norm (g x) = 1" if "x \<in> T" and "b \<bullet> x = 0" for x | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 520 | using that | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 521 | apply (simp only: g_def) | 
| 63130 | 522 | apply (rule power2_eq_imp_eq) | 
| 523 | apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps) | |
| 524 | apply (simp add: algebra_simps inner_commute) | |
| 525 | done | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 526 | have [simp]: "b \<bullet> g x \<noteq> 1" if "x \<in> T" and "b \<bullet> x = 0" for x | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 527 | using that unfolding g_def | 
| 63130 | 528 | apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff) | 
| 529 | apply (auto simp: algebra_simps) | |
| 530 | done | |
| 531 | have "subspace T" | |
| 532 | by (simp add: assms subspace_affine) | |
| 533 | have [simp]: "\<And>x. \<lbrakk>x \<in> T; b \<bullet> x = 0\<rbrakk> \<Longrightarrow> g x \<in> T" | |
| 534 | unfolding g_def | |
| 535 | by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff) | |
| 536 |   have "f ` {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<subseteq> {x. b\<bullet>x = 0}"
 | |
| 537 | unfolding f_def using \<open>norm b = 1\<close> norm_eq_1 | |
| 538 | by (force simp: field_simps inner_add_right inner_diff_right) | |
| 539 | moreover have "f ` T \<subseteq> T" | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70620diff
changeset | 540 | unfolding f_def using assms \<open>subspace T\<close> | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70620diff
changeset | 541 | by (auto simp add: inner_add_right inner_diff_right mem_affine_3_minus subspace_mul) | 
| 63130 | 542 |   moreover have "{x. b\<bullet>x = 0} \<inter> T \<subseteq> f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T)"
 | 
| 543 | apply clarify | |
| 544 | apply (rule_tac x = "g x" in image_eqI, auto) | |
| 545 | done | |
| 546 |   ultimately have imf: "f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) = {x. b\<bullet>x = 0} \<inter> T"
 | |
| 547 | by blast | |
| 548 | have no4: "\<And>y. b\<bullet>y = 0 \<Longrightarrow> norm ((y\<bullet>y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\<bullet>y + 4" | |
| 549 | apply (rule power2_eq_imp_eq) | |
| 550 | apply (simp_all add: dot_square_norm [symmetric]) | |
| 551 | apply (auto simp: power2_eq_square algebra_simps inner_commute) | |
| 552 | done | |
| 553 | have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> b \<bullet> f x = 0" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 554 | by (simp add: f_def algebra_simps field_split_simps) | 
| 63130 | 555 | have [simp]: "\<And>x. \<lbrakk>x \<in> T; norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> f x \<in> T" | 
| 556 | unfolding f_def | |
| 557 | by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff) | |
| 558 |   have "g ` {x. b\<bullet>x = 0} \<subseteq> {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1}"
 | |
| 559 | unfolding g_def | |
| 560 | apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric]) | |
| 561 | apply (auto simp: algebra_simps) | |
| 562 | done | |
| 563 | moreover have "g ` T \<subseteq> T" | |
| 564 | unfolding g_def | |
| 565 | by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff) | |
| 566 |   moreover have "{x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T \<subseteq> g ` ({x. b\<bullet>x = 0} \<inter> T)"
 | |
| 567 | apply clarify | |
| 568 | apply (rule_tac x = "f x" in image_eqI, auto) | |
| 569 | done | |
| 570 |   ultimately have img: "g ` ({x. b\<bullet>x = 0} \<inter> T) = {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T"
 | |
| 571 | by blast | |
| 572 |   have aff: "affine ({x. b\<bullet>x = 0} \<inter> T)"
 | |
| 573 | by (blast intro: affine_hyperplane assms) | |
| 574 |   have contf: "continuous_on ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) f"
 | |
| 575 | unfolding f_def by (rule continuous_intros | force)+ | |
| 576 |   have contg: "continuous_on ({x. b\<bullet>x = 0} \<inter> T) g"
 | |
| 577 | unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+ | |
| 578 |   have "(sphere 0 1 \<inter> T) - {b} = {x. norm x = 1 \<and> (b\<bullet>x \<noteq> 1)} \<inter> T"
 | |
| 579 | using \<open>norm b = 1\<close> by (auto simp: norm_eq_1) (metis vector_eq \<open>b\<bullet>b = 1\<close>) | |
| 580 |   also have "... homeomorphic {x. b\<bullet>x = 0} \<inter> T"
 | |
| 581 | by (rule homeomorphicI [OF imf img contf contg]) auto | |
| 582 | also have "... homeomorphic p" | |
| 583 | apply (rule homeomorphic_affine_sets [OF aff \<open>affine p\<close>]) | |
| 584 | apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \<open>affine T\<close>] affT) | |
| 585 | done | |
| 586 | finally show ?thesis . | |
| 587 | qed | |
| 588 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 589 | theorem homeomorphic_punctured_affine_sphere_affine: | 
| 63130 | 590 | fixes a :: "'a :: euclidean_space" | 
| 591 | assumes "0 < r" "b \<in> sphere a r" "affine T" "a \<in> T" "b \<in> T" "affine p" | |
| 592 | and aff: "aff_dim T = aff_dim p + 1" | |
| 66710 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 593 |     shows "(sphere a r \<inter> T) - {b} homeomorphic p"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 594 | proof - | 
| 63130 | 595 | have "a \<noteq> b" using assms by auto | 
| 596 | then have inj: "inj (\<lambda>x::'a. x /\<^sub>R norm (a - b))" | |
| 597 | by (simp add: inj_on_def) | |
| 598 |   have "((sphere a r \<inter> T) - {b}) homeomorphic
 | |
| 67399 | 599 |         (+) (-a) ` ((sphere a r \<inter> T) - {b})"
 | 
| 63130 | 600 | by (rule homeomorphic_translation) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 601 |   also have "... homeomorphic (*\<^sub>R) (inverse r) ` (+) (- a) ` (sphere a r \<inter> T - {b})"
 | 
| 63130 | 602 | by (metis \<open>0 < r\<close> homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 603 |   also have "... = sphere 0 1 \<inter> ((*\<^sub>R) (inverse r) ` (+) (- a) ` T) - {(b - a) /\<^sub>R r}"
 | 
| 63130 | 604 | using assms by (auto simp: dist_norm norm_minus_commute divide_simps) | 
| 605 | also have "... homeomorphic p" | |
| 606 | apply (rule homeomorphic_punctured_affine_sphere_affine_01) | |
| 69661 | 607 | using assms affine_translation [symmetric, of "- a"] aff_dim_translation_eq [of "- a"] | 
| 608 | apply (auto simp: dist_norm norm_minus_commute affine_scaling inj) | |
| 63130 | 609 | done | 
| 610 | finally show ?thesis . | |
| 611 | qed | |
| 612 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 613 | corollary homeomorphic_punctured_sphere_affine: | 
| 66710 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 614 | fixes a :: "'a :: euclidean_space" | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 615 | assumes "0 < r" and b: "b \<in> sphere a r" | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 616 |       and "affine T" and affS: "aff_dim T + 1 = DIM('a)"
 | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 617 |     shows "(sphere a r - {b}) homeomorphic T"
 | 
| 70136 | 618 | using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto | 
| 66710 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 619 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 620 | corollary homeomorphic_punctured_sphere_hyperplane: | 
| 66710 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 621 | fixes a :: "'a :: euclidean_space" | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 622 | assumes "0 < r" and b: "b \<in> sphere a r" | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 623 | and "c \<noteq> 0" | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 624 |     shows "(sphere a r - {b}) homeomorphic {x::'a. c \<bullet> x = d}"
 | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 625 | apply (rule homeomorphic_punctured_sphere_affine) | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 626 | using assms | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 627 | apply (auto simp: affine_hyperplane of_nat_diff) | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 628 | done | 
| 
676258a1cf01
eliminated a needless dependence on the theorem homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66690diff
changeset | 629 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 630 | proposition homeomorphic_punctured_sphere_affine_gen: | 
| 63130 | 631 | fixes a :: "'a :: euclidean_space" | 
| 632 | assumes "convex S" "bounded S" and a: "a \<in> rel_frontier S" | |
| 633 | and "affine T" and affS: "aff_dim S = aff_dim T + 1" | |
| 634 |     shows "rel_frontier S - {a} homeomorphic T"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 635 | proof - | 
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 636 | obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S" | 
| 63130 | 637 | using choose_affine_subset [OF affine_UNIV aff_dim_geq] | 
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 638 | by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex) | 
| 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 639 |   have "S \<noteq> {}" using assms by auto
 | 
| 63130 | 640 | then obtain z where "z \<in> U" | 
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 641 | by (metis aff_dim_negative_iff equals0I affdS) | 
| 63130 | 642 |   then have bne: "ball z 1 \<inter> U \<noteq> {}" by force
 | 
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 643 | then have [simp]: "aff_dim(ball z 1 \<inter> U) = aff_dim U" | 
| 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 644 | using aff_dim_convex_Int_open [OF \<open>convex U\<close> open_ball] | 
| 63130 | 645 | by (fastforce simp add: Int_commute) | 
| 646 | have "rel_frontier S homeomorphic rel_frontier (ball z 1 \<inter> U)" | |
| 68006 | 647 | by (rule homeomorphic_rel_frontiers_convex_bounded_sets) | 
| 648 | (auto simp: \<open>affine U\<close> affine_imp_convex convex_Int affdS assms) | |
| 63130 | 649 | also have "... = sphere z 1 \<inter> U" | 
| 650 | using convex_affine_rel_frontier_Int [of "ball z 1" U] | |
| 651 | by (simp add: \<open>affine U\<close> bne) | |
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 652 | finally have "rel_frontier S homeomorphic sphere z 1 \<inter> U" . | 
| 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 653 | then obtain h k where him: "h ` rel_frontier S = sphere z 1 \<inter> U" | 
| 63130 | 654 | and kim: "k ` (sphere z 1 \<inter> U) = rel_frontier S" | 
| 655 | and hcon: "continuous_on (rel_frontier S) h" | |
| 656 | and kcon: "continuous_on (sphere z 1 \<inter> U) k" | |
| 657 | and kh: "\<And>x. x \<in> rel_frontier S \<Longrightarrow> k(h(x)) = x" | |
| 658 | and hk: "\<And>y. y \<in> sphere z 1 \<inter> U \<Longrightarrow> h(k(y)) = y" | |
| 659 | unfolding homeomorphic_def homeomorphism_def by auto | |
| 660 |   have "rel_frontier S - {a} homeomorphic (sphere z 1 \<inter> U) - {h a}"
 | |
| 66690 
6953b1a29e19
Tiny presentational improvements to homeomorphic_punctured_sphere_affine_gen
 paulson <lp15@cam.ac.uk> parents: 
66287diff
changeset | 661 | proof (rule homeomorphicI) | 
| 63130 | 662 |     show h: "h ` (rel_frontier S - {a}) = sphere z 1 \<inter> U - {h a}"
 | 
| 663 | using him a kh by auto metis | |
| 664 |     show "k ` (sphere z 1 \<inter> U - {h a}) = rel_frontier S - {a}"
 | |
| 665 | by (force simp: h [symmetric] image_comp o_def kh) | |
| 666 | qed (auto intro: continuous_on_subset hcon kcon simp: kh hk) | |
| 667 | also have "... homeomorphic T" | |
| 68006 | 668 | by (rule homeomorphic_punctured_affine_sphere_affine) | 
| 669 | (use a him in \<open>auto simp: affS affdS \<open>affine T\<close> \<open>affine U\<close> \<open>z \<in> U\<close>\<close>) | |
| 63130 | 670 | finally show ?thesis . | 
| 671 | qed | |
| 672 | ||
| 673 | ||
| 674 | text\<open> When dealing with AR, ANR and ANR later, it's useful to know that every set | |
| 675 | is homeomorphic to a closed subset of a convex set, and | |
| 676 | if the set is locally compact we can take the convex set to be the universe.\<close> | |
| 677 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 678 | proposition homeomorphic_closedin_convex: | 
| 63130 | 679 | fixes S :: "'m::euclidean_space set" | 
| 680 |   assumes "aff_dim S < DIM('n)"
 | |
| 681 | obtains U and T :: "'n::euclidean_space set" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 682 |      where "convex U" "U \<noteq> {}" "closedin (top_of_set U) T"
 | 
| 63130 | 683 | "S homeomorphic T" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 684 | proof (cases "S = {}")
 | 
| 63130 | 685 | case True then show ?thesis | 
| 686 |     by (rule_tac U=UNIV and T="{}" in that) auto
 | |
| 687 | next | |
| 688 | case False | |
| 689 | then obtain a where "a \<in> S" by auto | |
| 690 | obtain i::'n where i: "i \<in> Basis" "i \<noteq> 0" | |
| 691 | using SOME_Basis Basis_zero by force | |
| 67399 | 692 | have "0 \<in> affine hull ((+) (- a) ` S)" | 
| 63130 | 693 | by (simp add: \<open>a \<in> S\<close> hull_inc) | 
| 67399 | 694 | then have "dim ((+) (- a) ` S) = aff_dim ((+) (- a) ` S)" | 
| 63130 | 695 | by (simp add: aff_dim_zero) | 
| 696 |   also have "... < DIM('n)"
 | |
| 69661 | 697 | by (simp add: aff_dim_translation_eq_subtract assms cong: image_cong_simp) | 
| 67399 | 698 |   finally have dd: "dim ((+) (- a) ` S) < DIM('n)"
 | 
| 63130 | 699 | by linarith | 
| 69661 | 700 |   have span: "span {x. i \<bullet> x = 0} = {x. i \<bullet> x = 0}"
 | 
| 701 |     using span_eq_iff [symmetric, of "{x. i \<bullet> x = 0}"] subspace_hyperplane [of i] by simp
 | |
| 702 |   have "dim ((+) (- a) ` S) \<le> dim {x. i \<bullet> x = 0}"
 | |
| 703 | using dd by (simp add: dim_hyperplane [OF \<open>i \<noteq> 0\<close>]) | |
| 704 |   then obtain T where "subspace T" and Tsub: "T \<subseteq> {x. i \<bullet> x = 0}"
 | |
| 705 | and dimT: "dim T = dim ((+) (- a) ` S)" | |
| 706 | by (rule choose_subspace_of_subspace) (simp add: span) | |
| 67399 | 707 | have "subspace (span ((+) (- a) ` S))" | 
| 63130 | 708 | using subspace_span by blast | 
| 709 | then obtain h k where "linear h" "linear k" | |
| 67399 | 710 | and heq: "h ` span ((+) (- a) ` S) = T" | 
| 711 | and keq:"k ` T = span ((+) (- a) ` S)" | |
| 712 | and hinv [simp]: "\<And>x. x \<in> span ((+) (- a) ` S) \<Longrightarrow> k(h x) = x" | |
| 63130 | 713 | and kinv [simp]: "\<And>x. x \<in> T \<Longrightarrow> h(k x) = x" | 
| 714 | apply (rule isometries_subspaces [OF _ \<open>subspace T\<close>]) | |
| 715 | apply (auto simp: dimT) | |
| 716 | done | |
| 717 | have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B | |
| 718 | using \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_conv_bounded_linear by blast+ | |
| 719 | have ihhhh[simp]: "\<And>x. x \<in> S \<Longrightarrow> i \<bullet> h (x - a) = 0" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67399diff
changeset | 720 | using Tsub [THEN subsetD] heq span_superset by fastforce | 
| 63130 | 721 |   have "sphere 0 1 - {i} homeomorphic {x. i \<bullet> x = 0}"
 | 
| 722 | apply (rule homeomorphic_punctured_sphere_affine) | |
| 723 | using i | |
| 724 | apply (auto simp: affine_hyperplane) | |
| 725 | by (metis DIM_positive Suc_eq_plus1 add.left_neutral diff_add_cancel not_le not_less_eq_eq of_nat_1 of_nat_diff) | |
| 726 |   then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \<bullet> x = 0} f g"
 | |
| 727 | by (force simp: homeomorphic_def) | |
| 67399 | 728 | have "h ` (+) (- a) ` S \<subseteq> T" | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68006diff
changeset | 729 | using heq span_superset span_linear_image by blast | 
| 67399 | 730 |   then have "g ` h ` (+) (- a) ` S \<subseteq> g ` {x. i \<bullet> x = 0}"
 | 
| 63130 | 731 | using Tsub by (simp add: image_mono) | 
| 732 |   also have "... \<subseteq> sphere 0 1 - {i}"
 | |
| 733 | by (simp add: fg [unfolded homeomorphism_def]) | |
| 67399 | 734 |   finally have gh_sub_sph: "(g \<circ> h) ` (+) (- a) ` S \<subseteq> sphere 0 1 - {i}"
 | 
| 63130 | 735 | by (metis image_comp) | 
| 67399 | 736 | then have gh_sub_cb: "(g \<circ> h) ` (+) (- a) ` S \<subseteq> cball 0 1" | 
| 63130 | 737 | by (metis Diff_subset order_trans sphere_cball) | 
| 738 | have [simp]: "\<And>u. u \<in> S \<Longrightarrow> norm (g (h (u - a))) = 1" | |
| 739 | using gh_sub_sph [THEN subsetD] by (auto simp: o_def) | |
| 69661 | 740 | have ghcont: "continuous_on ((\<lambda>x. x - a) ` S) (\<lambda>x. g (h x))" | 
| 63130 | 741 | apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force) | 
| 742 | done | |
| 69661 | 743 | have kfcont: "continuous_on ((\<lambda>x. g (h (x - a))) ` S) (\<lambda>x. k (f x))" | 
| 63130 | 744 | apply (rule continuous_on_compose2 [OF kcont]) | 
| 745 | using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast) | |
| 746 | done | |
| 67399 | 747 | have "S homeomorphic (+) (- a) ` S" | 
| 69661 | 748 | by (fact homeomorphic_translation) | 
| 749 | also have "\<dots> homeomorphic (g \<circ> h) ` (+) (- a) ` S" | |
| 750 | apply (simp add: homeomorphic_def homeomorphism_def cong: image_cong_simp) | |
| 63130 | 751 | apply (rule_tac x="g \<circ> h" in exI) | 
| 752 | apply (rule_tac x="k \<circ> f" in exI) | |
| 69661 | 753 | apply (auto simp: ghcont kfcont span_base homeomorphism_apply2 [OF fg] image_comp cong: image_cong_simp) | 
| 63130 | 754 | done | 
| 69661 | 755 | finally have Shom: "S homeomorphic (\<lambda>x. g (h x)) ` (\<lambda>x. x - a) ` S" | 
| 756 | by (simp cong: image_cong_simp) | |
| 63130 | 757 | show ?thesis | 
| 67399 | 758 | apply (rule_tac U = "ball 0 1 \<union> image (g o h) ((+) (- a) ` S)" | 
| 759 | and T = "image (g o h) ((+) (- a) ` S)" | |
| 63130 | 760 | in that) | 
| 761 | apply (rule convex_intermediate_ball [of 0 1], force) | |
| 762 | using gh_sub_cb apply force | |
| 763 | apply force | |
| 764 | apply (simp add: closedin_closed) | |
| 765 | apply (rule_tac x="sphere 0 1" in exI) | |
| 69661 | 766 | apply (auto simp: Shom cong: image_cong_simp) | 
| 63130 | 767 | done | 
| 768 | qed | |
| 769 | ||
| 69683 | 770 | subsection\<open>Locally compact sets in an open set\<close> | 
| 63130 | 771 | |
| 772 | text\<open> Locally compact sets are closed in an open set and are homeomorphic | |
| 773 | to an absolutely closed set if we have one more dimension to play with.\<close> | |
| 774 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 775 | lemma locally_compact_open_Int_closure: | 
| 63130 | 776 | fixes S :: "'a :: metric_space set" | 
| 777 | assumes "locally compact S" | |
| 778 | obtains T where "open T" "S = T \<inter> closure S" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 779 | proof - | 
| 63130 | 780 | have "\<forall>x\<in>S. \<exists>T v u. u = S \<inter> T \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> open T \<and> compact v" | 
| 781 | by (metis assms locally_compact openin_open) | |
| 782 | then obtain t v where | |
| 783 | tv: "\<And>x. x \<in> S | |
| 784 | \<Longrightarrow> v x \<subseteq> S \<and> open (t x) \<and> compact (v x) \<and> (\<exists>u. x \<in> u \<and> u \<subseteq> v x \<and> u = S \<inter> t x)" | |
| 785 | by metis | |
| 69313 | 786 | then have o: "open (\<Union>(t ` S))" | 
| 63130 | 787 | by blast | 
| 788 | have "S = \<Union> (v ` S)" | |
| 789 | using tv by blast | |
| 69313 | 790 | also have "... = \<Union>(t ` S) \<inter> closure S" | 
| 63130 | 791 | proof | 
| 69313 | 792 | show "\<Union>(v ` S) \<subseteq> \<Union>(t ` S) \<inter> closure S" | 
| 63130 | 793 | apply safe | 
| 794 | apply (metis Int_iff subsetD UN_iff tv) | |
| 795 | apply (simp add: closure_def rev_subsetD tv) | |
| 796 | done | |
| 797 | have "t x \<inter> closure S \<subseteq> v x" if "x \<in> S" for x | |
| 798 | proof - | |
| 799 | have "t x \<inter> closure S \<subseteq> closure (t x \<inter> S)" | |
| 800 | by (simp add: open_Int_closure_subset that tv) | |
| 801 | also have "... \<subseteq> v x" | |
| 802 | by (metis Int_commute closure_minimal compact_imp_closed that tv) | |
| 803 | finally show ?thesis . | |
| 804 | qed | |
| 69313 | 805 | then show "\<Union>(t ` S) \<inter> closure S \<subseteq> \<Union>(v ` S)" | 
| 63130 | 806 | by blast | 
| 807 | qed | |
| 69313 | 808 | finally have e: "S = \<Union>(t ` S) \<inter> closure S" . | 
| 63130 | 809 | show ?thesis | 
| 810 | by (rule that [OF o e]) | |
| 811 | qed | |
| 812 | ||
| 813 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 814 | lemma locally_compact_closedin_open: | 
| 63130 | 815 | fixes S :: "'a :: metric_space set" | 
| 816 | assumes "locally compact S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 817 | obtains T where "open T" "closedin (top_of_set T) S" | 
| 63130 | 818 | by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) | 
| 819 | ||
| 820 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 821 | lemma locally_compact_homeomorphism_projection_closed: | 
| 63130 | 822 | assumes "locally compact S" | 
| 823 | obtains T and f :: "'a \<Rightarrow> 'a :: euclidean_space \<times> 'b :: euclidean_space" | |
| 824 | where "closed T" "homeomorphism S T f fst" | |
| 825 | proof (cases "closed S") | |
| 826 | case True | |
| 827 | then show ?thesis | |
| 828 |       apply (rule_tac T = "S \<times> {0}" and f = "\<lambda>x. (x, 0)" in that)
 | |
| 829 | apply (auto simp: closed_Times homeomorphism_def continuous_intros) | |
| 830 | done | |
| 831 | next | |
| 832 | case False | |
| 833 | obtain U where "open U" and US: "U \<inter> closure S = S" | |
| 834 | by (metis locally_compact_open_Int_closure [OF assms]) | |
| 835 |     with False have Ucomp: "-U \<noteq> {}"
 | |
| 836 | using closure_eq by auto | |
| 837 | have [simp]: "closure (- U) = -U" | |
| 838 | by (simp add: \<open>open U\<close> closed_Compl) | |
| 839 |     define f :: "'a \<Rightarrow> 'a \<times> 'b" where "f \<equiv> \<lambda>x. (x, One /\<^sub>R setdist {x} (- U))"
 | |
| 840 |     have "continuous_on U (\<lambda>x. (x, One /\<^sub>R setdist {x} (- U)))"
 | |
| 63301 | 841 | apply (intro continuous_intros continuous_on_setdist) | 
| 842 | by (simp add: Ucomp setdist_eq_0_sing_1) | |
| 63130 | 843 | then have homU: "homeomorphism U (f`U) f fst" | 
| 844 | by (auto simp: f_def homeomorphism_def image_iff continuous_intros) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 845 | have cloS: "closedin (top_of_set U) S" | 
| 63130 | 846 | by (metis US closed_closure closedin_closed_Int) | 
| 847 |     have cont: "isCont ((\<lambda>x. setdist {x} (- U)) o fst) z" for z :: "'a \<times> 'b"
 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66710diff
changeset | 848 | by (rule continuous_at_compose continuous_intros continuous_at_setdist)+ | 
| 63130 | 849 |     have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \<Longrightarrow> setdist {a} (- U) \<noteq> 0" for a::'a and b::'b
 | 
| 850 | by force | |
| 851 | have *: "r *\<^sub>R b = One \<Longrightarrow> b = (1 / r) *\<^sub>R One" for r and b::'b | |
| 852 | by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one) | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 853 |     have "f ` U = (\<lambda>z. (setdist {fst z} (- U) *\<^sub>R snd z)) -` {One}"
 | 
| 63301 | 854 | apply (auto simp: f_def setdist_eq_0_sing_1 field_simps Ucomp) | 
| 63130 | 855 | apply (rule_tac x=a in image_eqI) | 
| 63301 | 856 | apply (auto simp: * setdist_eq_0_sing_1 dest: setdist1D) | 
| 63130 | 857 | done | 
| 858 | then have clfU: "closed (f ` U)" | |
| 859 | apply (rule ssubst) | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 860 | apply (rule continuous_closed_vimage) | 
| 63130 | 861 | apply (auto intro: continuous_intros cont [unfolded o_def]) | 
| 862 | done | |
| 863 | have "closed (f ` S)" | |
| 864 | apply (rule closedin_closed_trans [OF _ clfU]) | |
| 865 | apply (rule homeomorphism_imp_closed_map [OF homU cloS]) | |
| 866 | done | |
| 867 | then show ?thesis | |
| 868 | apply (rule that) | |
| 869 | apply (rule homeomorphism_of_subsets [OF homU]) | |
| 870 | using US apply auto | |
| 871 | done | |
| 872 | qed | |
| 873 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 874 | lemma locally_compact_closed_Int_open: | 
| 63130 | 875 | fixes S :: "'a :: euclidean_space set" | 
| 876 | shows | |
| 877 | "locally compact S \<longleftrightarrow> (\<exists>U u. closed U \<and> open u \<and> S = U \<inter> u)" | |
| 878 | by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact) | |
| 879 | ||
| 880 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 881 | lemma lowerdim_embeddings: | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 882 |   assumes  "DIM('a) < DIM('b)"
 | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 883 | obtains f :: "'a::euclidean_space*real \<Rightarrow> 'b::euclidean_space" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 884 | and g :: "'b \<Rightarrow> 'a*real" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 885 | and j :: 'b | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 886 | where "linear f" "linear g" "\<And>z. g (f z) = z" "j \<in> Basis" "\<And>x. f(x,0) \<bullet> j = 0" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 887 | proof - | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 888 |   let ?B = "Basis :: ('a*real) set"
 | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 889 | have b01: "(0,1) \<in> ?B" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 890 | by (simp add: Basis_prod_def) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 891 |   have "DIM('a * real) \<le> DIM('b)"
 | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 892 | by (simp add: Suc_leI assms) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 893 | then obtain basf :: "'a*real \<Rightarrow> 'b" where sbf: "basf ` ?B \<subseteq> Basis" and injbf: "inj_on basf Basis" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 894 | by (metis finite_Basis card_le_inj) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 895 | define basg:: "'b \<Rightarrow> 'a * real" where | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 896 | "basg \<equiv> \<lambda>i. if i \<in> basf ` Basis then inv_into Basis basf i else (0,1)" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 897 | have bgf[simp]: "basg (basf i) = i" if "i \<in> Basis" for i | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 898 | using inv_into_f_f injbf that by (force simp: basg_def) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 899 | have sbg: "basg ` Basis \<subseteq> ?B" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 900 | by (force simp: basg_def injbf b01) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 901 | define f :: "'a*real \<Rightarrow> 'b" where "f \<equiv> \<lambda>u. \<Sum>j\<in>Basis. (u \<bullet> basg j) *\<^sub>R j" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 902 | define g :: "'b \<Rightarrow> 'a*real" where "g \<equiv> \<lambda>z. (\<Sum>i\<in>Basis. (z \<bullet> basf i) *\<^sub>R i)" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 903 | show ?thesis | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 904 | proof | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 905 | show "linear f" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 906 | unfolding f_def | 
| 64267 | 907 | by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 908 | show "linear g" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 909 | unfolding g_def | 
| 64267 | 910 | by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 911 | have *: "(\<Sum>a \<in> Basis. a \<bullet> basf b * (x \<bullet> basg a)) = x \<bullet> b" if "b \<in> Basis" for x b | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 912 | using sbf that by auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 913 | show gf: "g (f x) = x" for x | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 914 | apply (rule euclidean_eqI) | 
| 64267 | 915 | apply (simp add: f_def g_def inner_sum_left scaleR_sum_left algebra_simps) | 
| 916 | apply (simp add: Groups_Big.sum_distrib_left [symmetric] *) | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 917 | done | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 918 | show "basf(0,1) \<in> Basis" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 919 | using b01 sbf by auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 920 | then show "f(x,0) \<bullet> basf(0,1) = 0" for x | 
| 64267 | 921 | apply (simp add: f_def inner_sum_left) | 
| 922 | apply (rule comm_monoid_add_class.sum.neutral) | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 923 | using b01 inner_not_same_Basis by fastforce | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 924 | qed | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 925 | qed | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 926 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 927 | proposition locally_compact_homeomorphic_closed: | 
| 63130 | 928 | fixes S :: "'a::euclidean_space set" | 
| 929 |   assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)"
 | |
| 930 | obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 931 | proof - | 
| 63130 | 932 |   obtain U:: "('a*real)set" and h
 | 
| 933 | where "closed U" and homU: "homeomorphism S U h fst" | |
| 934 | using locally_compact_homeomorphism_projection_closed assms by metis | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 935 | obtain f :: "'a*real \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'a*real" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 936 | where "linear f" "linear g" and gf [simp]: "\<And>z. g (f z) = z" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 937 | using lowerdim_embeddings [OF dimlt] by metis | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 938 | then have "inj f" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 939 | by (metis injI) | 
| 63130 | 940 | have gfU: "g ` f ` U = U" | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 941 | by (simp add: image_comp o_def) | 
| 63130 | 942 | have "S homeomorphic U" | 
| 943 | using homU homeomorphic_def by blast | |
| 944 | also have "... homeomorphic f ` U" | |
| 945 | apply (rule homeomorphicI [OF refl gfU]) | |
| 946 | apply (meson \<open>inj f\<close> \<open>linear f\<close> homeomorphism_cont2 linear_homeomorphism_image) | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 947 | using \<open>linear g\<close> linear_continuous_on linear_conv_bounded_linear apply blast | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 948 | apply (auto simp: o_def) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 949 | done | 
| 63130 | 950 | finally show ?thesis | 
| 951 | apply (rule_tac T = "f ` U" in that) | |
| 952 | apply (rule closed_injective_linear_image [OF \<open>closed U\<close> \<open>linear f\<close> \<open>inj f\<close>], assumption) | |
| 953 | done | |
| 954 | qed | |
| 955 | ||
| 956 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 957 | lemma homeomorphic_convex_compact_lemma: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 958 | fixes S :: "'a::euclidean_space set" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 959 | assumes "convex S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 960 | and "compact S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 961 | and "cball 0 1 \<subseteq> S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 962 | shows "S homeomorphic (cball (0::'a) 1)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 963 | proof (rule starlike_compact_projective_special[OF assms(2-3)]) | 
| 63130 | 964 | fix x u | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 965 | assume "x \<in> S" and "0 \<le> u" and "u < (1::real)" | 
| 63130 | 966 | have "open (ball (u *\<^sub>R x) (1 - u))" | 
| 967 | by (rule open_ball) | |
| 968 | moreover have "u *\<^sub>R x \<in> ball (u *\<^sub>R x) (1 - u)" | |
| 969 | unfolding centre_in_ball using \<open>u < 1\<close> by simp | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 970 | moreover have "ball (u *\<^sub>R x) (1 - u) \<subseteq> S" | 
| 63130 | 971 | proof | 
| 972 | fix y | |
| 973 | assume "y \<in> ball (u *\<^sub>R x) (1 - u)" | |
| 974 | then have "dist (u *\<^sub>R x) y < 1 - u" | |
| 975 | unfolding mem_ball . | |
| 976 | with \<open>u < 1\<close> have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1" | |
| 977 | by (simp add: dist_norm inverse_eq_divide norm_minus_commute) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 978 | with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> S" .. | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 979 | with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \<in> S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 980 | using \<open>x \<in> S\<close> \<open>0 \<le> u\<close> \<open>u < 1\<close> [THEN less_imp_le] by (rule convexD_alt) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 981 | then show "y \<in> S" using \<open>u < 1\<close> | 
| 63130 | 982 | by simp | 
| 983 | qed | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 984 | ultimately have "u *\<^sub>R x \<in> interior S" .. | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 985 | then show "u *\<^sub>R x \<in> S - frontier S" | 
| 63130 | 986 | using frontier_def and interior_subset by auto | 
| 987 | qed | |
| 988 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 989 | proposition homeomorphic_convex_compact_cball: | 
| 63130 | 990 | fixes e :: real | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 991 | and S :: "'a::euclidean_space set" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 992 | assumes "convex S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 993 | and "compact S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 994 |     and "interior S \<noteq> {}"
 | 
| 63130 | 995 | and "e > 0" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 996 | shows "S homeomorphic (cball (b::'a) e)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 997 | proof - | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 998 | obtain a where "a \<in> interior S" | 
| 63130 | 999 | using assms(3) by auto | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1000 | then obtain d where "d > 0" and d: "cball a d \<subseteq> S" | 
| 63130 | 1001 | unfolding mem_interior_cball by auto | 
| 1002 | let ?d = "inverse d" and ?n = "0::'a" | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1003 | have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` S" | 
| 63130 | 1004 | apply rule | 
| 1005 | apply (rule_tac x="d *\<^sub>R x + a" in image_eqI) | |
| 1006 | defer | |
| 1007 | apply (rule d[unfolded subset_eq, rule_format]) | |
| 1008 | using \<open>d > 0\<close> | |
| 1009 | unfolding mem_cball dist_norm | |
| 1010 | apply (auto simp add: mult_right_le_one_le) | |
| 1011 | done | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1012 | then have "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` S homeomorphic cball ?n 1" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1013 | using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` S", | 
| 63130 | 1014 | OF convex_affinity compact_affinity] | 
| 1015 | using assms(1,2) | |
| 1016 | by (auto simp add: scaleR_right_diff_distrib) | |
| 1017 | then show ?thesis | |
| 1018 | apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]]) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1019 | apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" S "?d *\<^sub>R -a"]]) | 
| 63130 | 1020 | using \<open>d>0\<close> \<open>e>0\<close> | 
| 1021 | apply (auto simp add: scaleR_right_diff_distrib) | |
| 1022 | done | |
| 1023 | qed | |
| 1024 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1025 | corollary homeomorphic_convex_compact: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1026 | fixes S :: "'a::euclidean_space set" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1027 | and T :: "'a set" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1028 |   assumes "convex S" "compact S" "interior S \<noteq> {}"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1029 |     and "convex T" "compact T" "interior T \<noteq> {}"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1030 | shows "S homeomorphic T" | 
| 63130 | 1031 | using assms | 
| 1032 | by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) | |
| 1033 | ||
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1034 | lemma homeomorphic_closed_intervals: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1035 | fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1036 |   assumes "box a b \<noteq> {}" and "box c d \<noteq> {}"
 | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1037 | shows "(cbox a b) homeomorphic (cbox c d)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1038 | apply (rule homeomorphic_convex_compact) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1039 | using assms apply auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1040 | done | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1041 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1042 | lemma homeomorphic_closed_intervals_real: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1043 | fixes a::real and b and c::real and d | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1044 | assumes "a<b" and "c<d" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1045 |   shows "{a..b} homeomorphic {c..d}"
 | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1046 | using assms by (auto intro: homeomorphic_convex_compact) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 1047 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1048 | subsection\<open>Covering spaces and lifting results for them\<close> | 
| 63301 | 1049 | |
| 70136 | 1050 | definition\<^marker>\<open>tag important\<close> covering_space | 
| 63301 | 1051 |            :: "'a::topological_space set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
 | 
| 1052 | where | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1053 | "covering_space c p S \<equiv> | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1054 | continuous_on c p \<and> p ` c = S \<and> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1055 | (\<forall>x \<in> S. \<exists>T. x \<in> T \<and> openin (top_of_set S) T \<and> | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1056 | (\<exists>v. \<Union>v = c \<inter> p -` T \<and> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1057 | (\<forall>u \<in> v. openin (top_of_set c) u) \<and> | 
| 63301 | 1058 | pairwise disjnt v \<and> | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1059 | (\<forall>u \<in> v. \<exists>q. homeomorphism u T p q)))" | 
| 63301 | 1060 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1061 | lemma covering_space_imp_continuous: "covering_space c p S \<Longrightarrow> continuous_on c p" | 
| 63301 | 1062 | by (simp add: covering_space_def) | 
| 1063 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1064 | lemma covering_space_imp_surjective: "covering_space c p S \<Longrightarrow> p ` c = S" | 
| 63301 | 1065 | by (simp add: covering_space_def) | 
| 1066 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1067 | lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \<Longrightarrow> covering_space S f T" | 
| 63301 | 1068 | apply (simp add: homeomorphism_def covering_space_def, clarify) | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1069 | apply (rule_tac x=T in exI, simp) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1070 |   apply (rule_tac x="{S}" in exI, auto)
 | 
| 63301 | 1071 | done | 
| 1072 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1073 | lemma covering_space_local_homeomorphism: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1074 | assumes "covering_space c p S" "x \<in> c" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1075 | obtains T u q where "x \<in> T" "openin (top_of_set c) T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1076 | "p x \<in> u" "openin (top_of_set S) u" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1077 | "homeomorphism T u p q" | 
| 63301 | 1078 | using assms | 
| 1079 | apply (simp add: covering_space_def, clarify) | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1080 | apply (drule_tac x="p x" in bspec, force) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1081 | by (metis IntI UnionE vimage_eq) | 
| 63301 | 1082 | |
| 1083 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1084 | lemma covering_space_local_homeomorphism_alt: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1085 | assumes p: "covering_space c p S" and "y \<in> S" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1086 | obtains x T U q where "p x = y" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1087 | "x \<in> T" "openin (top_of_set c) T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1088 | "y \<in> U" "openin (top_of_set S) U" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1089 | "homeomorphism T U p q" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1090 | proof - | 
| 63301 | 1091 | obtain x where "p x = y" "x \<in> c" | 
| 1092 | using assms covering_space_imp_surjective by blast | |
| 1093 | show ?thesis | |
| 1094 | apply (rule covering_space_local_homeomorphism [OF p \<open>x \<in> c\<close>]) | |
| 1095 | using that \<open>p x = y\<close> by blast | |
| 1096 | qed | |
| 1097 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1098 | proposition covering_space_open_map: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1099 | fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1100 | assumes p: "covering_space c p S" and T: "openin (top_of_set c) T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1101 | shows "openin (top_of_set S) (p ` T)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1102 | proof - | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1103 | have pce: "p ` c = S" | 
| 63301 | 1104 | and covs: | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1105 | "\<And>x. x \<in> S \<Longrightarrow> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1106 | \<exists>X VS. x \<in> X \<and> openin (top_of_set S) X \<and> | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1107 | \<Union>VS = c \<inter> p -` X \<and> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1108 | (\<forall>u \<in> VS. openin (top_of_set c) u) \<and> | 
| 63301 | 1109 | pairwise disjnt VS \<and> | 
| 1110 | (\<forall>u \<in> VS. \<exists>q. homeomorphism u X p q)" | |
| 1111 | using p by (auto simp: covering_space_def) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1112 | have "T \<subseteq> c" by (metis openin_euclidean_subtopology_iff T) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1113 | have "\<exists>X. openin (top_of_set S) X \<and> y \<in> X \<and> X \<subseteq> p ` T" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1114 | if "y \<in> p ` T" for y | 
| 63301 | 1115 | proof - | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1116 | have "y \<in> S" using \<open>T \<subseteq> c\<close> pce that by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1117 | obtain U VS where "y \<in> U" and U: "openin (top_of_set S) U" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1118 | and VS: "\<Union>VS = c \<inter> p -` U" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1119 | and openVS: "\<forall>V \<in> VS. openin (top_of_set c) V" | 
| 63301 | 1120 | and homVS: "\<And>V. V \<in> VS \<Longrightarrow> \<exists>q. homeomorphism V U p q" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1121 | using covs [OF \<open>y \<in> S\<close>] by auto | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1122 | obtain x where "x \<in> c" "p x \<in> U" "x \<in> T" "p x = y" | 
| 63301 | 1123 | apply simp | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1124 | using T [unfolded openin_euclidean_subtopology_iff] \<open>y \<in> U\<close> \<open>y \<in> p ` T\<close> by blast | 
| 63301 | 1125 | with VS obtain V where "x \<in> V" "V \<in> VS" by auto | 
| 1126 | then obtain q where q: "homeomorphism V U p q" using homVS by blast | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1127 | then have ptV: "p ` (T \<inter> V) = U \<inter> q -` (T \<inter> V)" | 
| 63301 | 1128 | using VS \<open>V \<in> VS\<close> by (auto simp: homeomorphism_def) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1129 | have ocv: "openin (top_of_set c) V" | 
| 63301 | 1130 | by (simp add: \<open>V \<in> VS\<close> openVS) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1131 | have "openin (top_of_set U) (U \<inter> q -` (T \<inter> V))" | 
| 63301 | 1132 | apply (rule continuous_on_open [THEN iffD1, rule_format]) | 
| 1133 | using homeomorphism_def q apply blast | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1134 | using openin_subtopology_Int_subset [of c] q T unfolding homeomorphism_def | 
| 63301 | 1135 | by (metis inf.absorb_iff2 Int_commute ocv openin_euclidean_subtopology_iff) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1136 | then have os: "openin (top_of_set S) (U \<inter> q -` (T \<inter> V))" | 
| 63301 | 1137 | using openin_trans [of U] by (simp add: Collect_conj_eq U) | 
| 1138 | show ?thesis | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1139 | apply (rule_tac x = "p ` (T \<inter> V)" in exI) | 
| 63301 | 1140 | apply (rule conjI) | 
| 1141 | apply (simp only: ptV os) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1142 | using \<open>p x = y\<close> \<open>x \<in> V\<close> \<open>x \<in> T\<close> apply blast | 
| 63301 | 1143 | done | 
| 1144 | qed | |
| 1145 | with openin_subopen show ?thesis by blast | |
| 1146 | qed | |
| 1147 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1148 | lemma covering_space_lift_unique_gen: | 
| 63301 | 1149 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 1150 | fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector" | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1151 | assumes cov: "covering_space c p S" | 
| 63301 | 1152 | and eq: "g1 a = g2 a" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1153 | and f: "continuous_on T f" "f ` T \<subseteq> S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1154 | and g1: "continuous_on T g1" "g1 ` T \<subseteq> c" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1155 | and fg1: "\<And>x. x \<in> T \<Longrightarrow> f x = p(g1 x)" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1156 | and g2: "continuous_on T g2" "g2 ` T \<subseteq> c" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1157 | and fg2: "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1158 | and u_compt: "U \<in> components T" and "a \<in> U" "x \<in> U" | 
| 63301 | 1159 | shows "g1 x = g2 x" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1160 | proof - | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1161 | have "U \<subseteq> T" by (rule in_components_subset [OF u_compt]) | 
| 65064 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 paulson <lp15@cam.ac.uk> parents: 
64792diff
changeset | 1162 |   define G12 where "G12 \<equiv> {x \<in> U. g1 x - g2 x = 0}"
 | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1163 | have "connected U" by (rule in_components_connected [OF u_compt]) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1164 | have contu: "continuous_on U g1" "continuous_on U g2" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1165 | using \<open>U \<subseteq> T\<close> continuous_on_subset g1 g2 by blast+ | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1166 | have o12: "openin (top_of_set U) G12" | 
| 63301 | 1167 | unfolding G12_def | 
| 1168 | proof (subst openin_subopen, clarify) | |
| 1169 | fix z | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1170 | assume z: "z \<in> U" "g1 z - g2 z = 0" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1171 | obtain v w q where "g1 z \<in> v" and ocv: "openin (top_of_set c) v" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1172 | and "p (g1 z) \<in> w" and osw: "openin (top_of_set S) w" | 
| 63301 | 1173 | and hom: "homeomorphism v w p q" | 
| 1174 | apply (rule_tac x = "g1 z" in covering_space_local_homeomorphism [OF cov]) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1175 | using \<open>U \<subseteq> T\<close> \<open>z \<in> U\<close> g1(2) apply blast+ | 
| 63301 | 1176 | done | 
| 1177 | have "g2 z \<in> v" using \<open>g1 z \<in> v\<close> z by auto | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1178 | have gg: "U \<inter> g -` v = U \<inter> g -` (v \<inter> g ` U)" for g | 
| 63301 | 1179 | by auto | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1180 | have "openin (top_of_set (g1 ` U)) (v \<inter> g1 ` U)" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1181 | using ocv \<open>U \<subseteq> T\<close> g1 by (fastforce simp add: openin_open) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1182 | then have 1: "openin (top_of_set U) (U \<inter> g1 -` v)" | 
| 63301 | 1183 | unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1184 | have "openin (top_of_set (g2 ` U)) (v \<inter> g2 ` U)" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1185 | using ocv \<open>U \<subseteq> T\<close> g2 by (fastforce simp add: openin_open) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1186 | then have 2: "openin (top_of_set U) (U \<inter> g2 -` v)" | 
| 63301 | 1187 | unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1188 |     show "\<exists>T. openin (top_of_set U) T \<and> z \<in> T \<and> T \<subseteq> {z \<in> U. g1 z - g2 z = 0}"
 | 
| 63301 | 1189 | using z | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1190 | apply (rule_tac x = "(U \<inter> g1 -` v) \<inter> (U \<inter> g2 -` v)" in exI) | 
| 63301 | 1191 | apply (intro conjI) | 
| 1192 | apply (rule openin_Int [OF 1 2]) | |
| 1193 | using \<open>g1 z \<in> v\<close> \<open>g2 z \<in> v\<close> apply (force simp:, clarify) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1194 | apply (metis \<open>U \<subseteq> T\<close> subsetD eq_iff_diff_eq_0 fg1 fg2 hom homeomorphism_def) | 
| 63301 | 1195 | done | 
| 1196 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1197 | have c12: "closedin (top_of_set U) G12" | 
| 63301 | 1198 | unfolding G12_def | 
| 1199 | by (intro continuous_intros continuous_closedin_preimage_constant contu) | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1200 |   have "G12 = {} \<or> G12 = U"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1201 | by (intro connected_clopen [THEN iffD1, rule_format] \<open>connected U\<close> conjI o12 c12) | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1202 | with eq \<open>a \<in> U\<close> have "\<And>x. x \<in> U \<Longrightarrow> g1 x - g2 x = 0" by (auto simp: G12_def) | 
| 63301 | 1203 | then show ?thesis | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1204 | using \<open>x \<in> U\<close> by force | 
| 63301 | 1205 | qed | 
| 1206 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1207 | proposition covering_space_lift_unique: | 
| 63301 | 1208 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 1209 | fixes g1 :: "'a \<Rightarrow> 'c::real_normed_vector" | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1210 | assumes "covering_space c p S" | 
| 63301 | 1211 | "g1 a = g2 a" | 
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1212 | "continuous_on T f" "f ` T \<subseteq> S" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1213 | "continuous_on T g1" "g1 ` T \<subseteq> c" "\<And>x. x \<in> T \<Longrightarrow> f x = p(g1 x)" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1214 | "continuous_on T g2" "g2 ` T \<subseteq> c" "\<And>x. x \<in> T \<Longrightarrow> f x = p(g2 x)" | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1215 | "connected T" "a \<in> T" "x \<in> T" | 
| 63301 | 1216 | shows "g1 x = g2 x" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1217 | using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv | 
| 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1218 | by blast | 
| 63301 | 1219 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1220 | lemma covering_space_locally: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1221 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1222 | assumes loc: "locally \<phi> C" and cov: "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1223 | and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1224 | shows "locally \<psi> S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1225 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1226 | have "locally \<psi> (p ` C)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1227 | apply (rule locally_open_map_image [OF loc]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1228 | using cov covering_space_imp_continuous apply blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1229 | using cov covering_space_imp_surjective covering_space_open_map apply blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1230 | by (simp add: pim) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1231 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1232 | using covering_space_imp_surjective [OF cov] by metis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1233 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1234 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1235 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1236 | proposition covering_space_locally_eq: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1237 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1238 | assumes cov: "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1239 | and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1240 | and qim: "\<And>q U. \<lbrakk>U \<subseteq> S; continuous_on U q; \<psi> U\<rbrakk> \<Longrightarrow> \<phi>(q ` U)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1241 | shows "locally \<psi> S \<longleftrightarrow> locally \<phi> C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1242 | (is "?lhs = ?rhs") | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1243 | proof | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1244 | assume L: ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1245 | show ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1246 | proof (rule locallyI) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1247 | fix V x | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1248 | assume V: "openin (top_of_set C) V" and "x \<in> V" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1249 | have "p x \<in> p ` C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1250 | by (metis IntE V \<open>x \<in> V\<close> imageI openin_open) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1251 | then obtain T \<V> where "p x \<in> T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1252 | and opeT: "openin (top_of_set S) T" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1253 | and veq: "\<Union>\<V> = C \<inter> p -` T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1254 | and ope: "\<forall>U\<in>\<V>. openin (top_of_set C) U" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1255 | and hom: "\<forall>U\<in>\<V>. \<exists>q. homeomorphism U T p q" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1256 | using cov unfolding covering_space_def by (blast intro: that) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1257 | have "x \<in> \<Union>\<V>" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1258 | using V veq \<open>p x \<in> T\<close> \<open>x \<in> V\<close> openin_imp_subset by fastforce | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1259 | then obtain U where "x \<in> U" "U \<in> \<V>" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1260 | by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1261 | then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1262 | using ope hom by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1263 | with V have "openin (top_of_set C) (U \<inter> V)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1264 | by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1265 | then have UV: "openin (top_of_set S) (p ` (U \<inter> V))" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1266 | using cov covering_space_open_map by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1267 | obtain W W' where opeW: "openin (top_of_set S) W" and "\<psi> W'" "p x \<in> W" "W \<subseteq> W'" and W'sub: "W' \<subseteq> p ` (U \<inter> V)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1268 | using locallyE [OF L UV] \<open>x \<in> U\<close> \<open>x \<in> V\<close> by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1269 | then have "W \<subseteq> T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1270 | by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1271 | show "\<exists>U Z. openin (top_of_set C) U \<and> | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1272 | \<phi> Z \<and> x \<in> U \<and> U \<subseteq> Z \<and> Z \<subseteq> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1273 | proof (intro exI conjI) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1274 | have "openin (top_of_set T) W" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1275 | by (meson opeW opeT openin_imp_subset openin_subset_trans \<open>W \<subseteq> T\<close>) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1276 | then have "openin (top_of_set U) (q ` W)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1277 | by (meson homeomorphism_imp_open_map homeomorphism_symD q) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1278 | then show "openin (top_of_set C) (q ` W)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1279 | using opeU openin_trans by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1280 | show "\<phi> (q ` W')" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1281 | by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \<open>\<psi> W'\<close> continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1282 | show "x \<in> q ` W" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1283 | by (metis \<open>p x \<in> W\<close> \<open>x \<in> U\<close> homeomorphism_def imageI q) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1284 | show "q ` W \<subseteq> q ` W'" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1285 | using \<open>W \<subseteq> W'\<close> by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1286 | have "W' \<subseteq> p ` V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1287 | using W'sub by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1288 | then show "q ` W' \<subseteq> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1289 | using W'sub homeomorphism_apply1 [OF q] by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1290 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1291 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1292 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1293 | assume ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1294 | then show ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1295 | using cov covering_space_locally pim by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1296 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1297 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1298 | lemma covering_space_locally_compact_eq: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1299 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1300 | assumes "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1301 | shows "locally compact S \<longleftrightarrow> locally compact C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1302 | apply (rule covering_space_locally_eq [OF assms]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1303 | apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1304 | using compact_continuous_image by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1305 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1306 | lemma covering_space_locally_connected_eq: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1307 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1308 | assumes "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1309 | shows "locally connected S \<longleftrightarrow> locally connected C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1310 | apply (rule covering_space_locally_eq [OF assms]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1311 | apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1312 | using connected_continuous_image by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1313 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1314 | lemma covering_space_locally_path_connected_eq: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1315 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1316 | assumes "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1317 | shows "locally path_connected S \<longleftrightarrow> locally path_connected C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1318 | apply (rule covering_space_locally_eq [OF assms]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1319 | apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1320 | using path_connected_continuous_image by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1321 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1322 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1323 | lemma covering_space_locally_compact: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1324 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1325 | assumes "locally compact C" "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1326 | shows "locally compact S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1327 | using assms covering_space_locally_compact_eq by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1328 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1329 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1330 | lemma covering_space_locally_connected: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1331 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1332 | assumes "locally connected C" "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1333 | shows "locally connected S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1334 | using assms covering_space_locally_connected_eq by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1335 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1336 | lemma covering_space_locally_path_connected: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1337 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1338 | assumes "locally path_connected C" "covering_space C p S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1339 | shows "locally path_connected S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1340 | using assms covering_space_locally_path_connected_eq by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 1341 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1342 | proposition covering_space_lift_homotopy: | 
| 64792 | 1343 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1344 | and h :: "real \<times> 'c::real_normed_vector \<Rightarrow> 'b" | |
| 1345 | assumes cov: "covering_space C p S" | |
| 1346 |       and conth: "continuous_on ({0..1} \<times> U) h"
 | |
| 1347 |       and him: "h ` ({0..1} \<times> U) \<subseteq> S"
 | |
| 1348 | and heq: "\<And>y. y \<in> U \<Longrightarrow> h (0,y) = p(f y)" | |
| 1349 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C" | |
| 1350 |     obtains k where "continuous_on ({0..1} \<times> U) k"
 | |
| 1351 |                     "k ` ({0..1} \<times> U) \<subseteq> C"
 | |
| 1352 | "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y" | |
| 1353 |                     "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1354 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1355 | have "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and> | 
| 64792 | 1356 |               continuous_on ({0..1} \<times> V) k \<and> k ` ({0..1} \<times> V) \<subseteq> C \<and>
 | 
| 1357 |               (\<forall>z \<in> V. k(0, z) = f z) \<and> (\<forall>z \<in> {0..1} \<times> V. h z = p(k z))"
 | |
| 1358 | if "y \<in> U" for y | |
| 1359 | proof - | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1360 | obtain UU where UU: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s) \<and> | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1361 | (\<exists>\<V>. \<Union>\<V> = C \<inter> p -` UU s \<and> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1362 | (\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and> | 
| 64792 | 1363 | pairwise disjnt \<V> \<and> | 
| 1364 | (\<forall>U \<in> \<V>. \<exists>q. homeomorphism U (UU s) p q))" | |
| 1365 | using cov unfolding covering_space_def by (metis (mono_tags)) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1366 | then have ope: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s)" | 
| 64792 | 1367 | by blast | 
| 1368 | have "\<exists>k n i. open k \<and> open n \<and> | |
| 1369 |                   t \<in> k \<and> y \<in> n \<and> i \<in> S \<and> h ` (({0..1} \<inter> k) \<times> (U \<inter> n)) \<subseteq> UU i" if "t \<in> {0..1}" for t
 | |
| 1370 | proof - | |
| 1371 | have hinS: "h (t, y) \<in> S" | |
| 1372 | using \<open>y \<in> U\<close> him that by blast | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1373 |       then have "(t,y) \<in> ({0..1} \<times> U) \<inter> h -` UU(h(t, y))"
 | 
| 64792 | 1374 |         using \<open>y \<in> U\<close> \<open>t \<in> {0..1}\<close>  by (auto simp: ope)
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1375 |       moreover have ope_01U: "openin (top_of_set ({0..1} \<times> U)) (({0..1} \<times> U) \<inter> h -` UU(h(t, y)))"
 | 
| 64792 | 1376 | using hinS ope continuous_on_open_gen [OF him] conth by blast | 
| 1377 |       ultimately obtain V W where opeV: "open V" and "t \<in> {0..1} \<inter> V" "t \<in> {0..1} \<inter> V"
 | |
| 1378 | and opeW: "open W" and "y \<in> U" "y \<in> W" | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1379 |                               and VW: "({0..1} \<inter> V) \<times> (U \<inter> W)  \<subseteq> (({0..1} \<times> U) \<inter> h -` UU(h(t, y)))"
 | 
| 64792 | 1380 | by (rule Times_in_interior_subtopology) (auto simp: openin_open) | 
| 1381 | then show ?thesis | |
| 1382 | using hinS by blast | |
| 1383 | qed | |
| 1384 | then obtain K NN X where | |
| 1385 |               K: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (K t)"
 | |
| 1386 |           and NN: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (NN t)"
 | |
| 1387 |           and inUS: "\<And>t. t \<in> {0..1} \<Longrightarrow> t \<in> K t \<and> y \<in> NN t \<and> X t \<in> S"
 | |
| 1388 |           and him: "\<And>t. t \<in> {0..1} \<Longrightarrow> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t)) \<subseteq> UU (X t)"
 | |
| 1389 | by (metis (mono_tags)) | |
| 1390 |     obtain \<T> where "\<T> \<subseteq> ((\<lambda>i. K i \<times> NN i)) ` {0..1}" "finite \<T>" "{0::real..1} \<times> {y} \<subseteq> \<Union>\<T>"
 | |
| 1391 | proof (rule compactE) | |
| 1392 |       show "compact ({0::real..1} \<times> {y})"
 | |
| 1393 | by (simp add: compact_Times) | |
| 1394 |       show "{0..1} \<times> {y} \<subseteq> (\<Union>i\<in>{0..1}. K i \<times> NN i)"
 | |
| 1395 | using K inUS by auto | |
| 1396 |       show "\<And>B. B \<in> (\<lambda>i. K i \<times> NN i) ` {0..1} \<Longrightarrow> open B"
 | |
| 1397 | using K NN by (auto simp: open_Times) | |
| 1398 | qed blast | |
| 1399 |     then obtain tk where "tk \<subseteq> {0..1}" "finite tk"
 | |
| 1400 |                      and tk: "{0::real..1} \<times> {y} \<subseteq> (\<Union>i \<in> tk. K i \<times> NN i)"
 | |
| 1401 | by (metis (no_types, lifting) finite_subset_image) | |
| 1402 |     then have "tk \<noteq> {}"
 | |
| 1403 | by auto | |
| 69313 | 1404 | define n where "n = \<Inter>(NN ` tk)" | 
| 64792 | 1405 | have "y \<in> n" "open n" | 
| 1406 |       using inUS NN \<open>tk \<subseteq> {0..1}\<close> \<open>finite tk\<close>
 | |
| 1407 | by (auto simp: n_def open_INT subset_iff) | |
| 1408 |     obtain \<delta> where "0 < \<delta>" and \<delta>: "\<And>T. \<lbrakk>T \<subseteq> {0..1}; diameter T < \<delta>\<rbrakk> \<Longrightarrow> \<exists>B\<in>K ` tk. T \<subseteq> B"
 | |
| 1409 |     proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"])
 | |
| 1410 |       show "K ` tk \<noteq> {}"
 | |
| 1411 |         using \<open>tk \<noteq> {}\<close> by auto
 | |
| 69313 | 1412 |       show "{0..1} \<subseteq> \<Union>(K ` tk)"
 | 
| 64792 | 1413 | using tk by auto | 
| 1414 | show "\<And>B. B \<in> K ` tk \<Longrightarrow> open B" | |
| 1415 |         using \<open>tk \<subseteq> {0..1}\<close> K by auto
 | |
| 1416 | qed auto | |
| 1417 | obtain N::nat where N: "N > 1 / \<delta>" | |
| 1418 | using reals_Archimedean2 by blast | |
| 1419 | then have "N > 0" | |
| 1420 | using \<open>0 < \<delta>\<close> order.asym by force | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1421 | have *: "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and> | 
| 64792 | 1422 |                    continuous_on ({0..of_nat n / N} \<times> V) k \<and>
 | 
| 1423 |                    k ` ({0..of_nat n / N} \<times> V) \<subseteq> C \<and>
 | |
| 1424 | (\<forall>z\<in>V. k (0, z) = f z) \<and> | |
| 1425 |                    (\<forall>z\<in>{0..of_nat n / N} \<times> V. h z = p (k z))" if "n \<le> N" for n
 | |
| 1426 | using that | |
| 1427 | proof (induction n) | |
| 1428 | case 0 | |
| 1429 | show ?case | |
| 1430 | apply (rule_tac x=U in exI) | |
| 1431 | apply (rule_tac x="f \<circ> snd" in exI) | |
| 1432 | apply (intro conjI \<open>y \<in> U\<close> continuous_intros continuous_on_subset [OF contf]) | |
| 1433 | using fim apply (auto simp: heq) | |
| 1434 | done | |
| 1435 | next | |
| 1436 | case (Suc n) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1437 | then obtain V k where opeUV: "openin (top_of_set U) V" | 
| 64792 | 1438 | and "y \<in> V" | 
| 1439 |                         and contk: "continuous_on ({0..real n / real N} \<times> V) k"
 | |
| 1440 |                         and kim: "k ` ({0..real n / real N} \<times> V) \<subseteq> C"
 | |
| 1441 | and keq: "\<And>z. z \<in> V \<Longrightarrow> k (0, z) = f z" | |
| 1442 |                         and heq: "\<And>z. z \<in> {0..real n / real N} \<times> V \<Longrightarrow> h z = p (k z)"
 | |
| 1443 | using Suc_leD by auto | |
| 1444 | have "n \<le> N" | |
| 1445 | using Suc.prems by auto | |
| 1446 |       obtain t where "t \<in> tk" and t: "{real n / real N .. (1 + real n) / real N} \<subseteq> K t"
 | |
| 1447 | proof (rule bexE [OF \<delta>]) | |
| 1448 |         show "{real n / real N .. (1 + real n) / real N} \<subseteq> {0..1}"
 | |
| 71172 | 1449 | using Suc.prems by (auto simp: field_split_simps) | 
| 64792 | 1450 |         show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \<delta>"
 | 
| 71172 | 1451 | using \<open>0 < \<delta>\<close> N by (auto simp: field_split_simps) | 
| 64792 | 1452 | qed blast | 
| 1453 |       have t01: "t \<in> {0..1}"
 | |
| 1454 |         using \<open>t \<in> tk\<close> \<open>tk \<subseteq> {0..1}\<close> by blast
 | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1455 | obtain \<V> where \<V>: "\<Union>\<V> = C \<inter> p -` UU (X t)" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1456 | and opeC: "\<And>U. U \<in> \<V> \<Longrightarrow> openin (top_of_set C) U" | 
| 64792 | 1457 | and "pairwise disjnt \<V>" | 
| 1458 | and homuu: "\<And>U. U \<in> \<V> \<Longrightarrow> \<exists>q. homeomorphism U (UU (X t)) p q" | |
| 1459 | using inUS [OF t01] UU by meson | |
| 1460 |       have n_div_N_in: "real n / real N \<in> {real n / real N .. (1 + real n) / real N}"
 | |
| 71172 | 1461 | using N by (auto simp: field_split_simps) | 
| 64792 | 1462 | with t have nN_in_kkt: "real n / real N \<in> K t" | 
| 1463 | by blast | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1464 | have "k (real n / real N, y) \<in> C \<inter> p -` UU (X t)" | 
| 64792 | 1465 | proof (simp, rule conjI) | 
| 1466 | show "k (real n / real N, y) \<in> C" | |
| 1467 | using \<open>y \<in> V\<close> kim keq by force | |
| 1468 | have "p (k (real n / real N, y)) = h (real n / real N, y)" | |
| 1469 | by (simp add: \<open>y \<in> V\<close> heq) | |
| 1470 |         also have "... \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
 | |
| 1471 | apply (rule imageI) | |
| 1472 | using \<open>y \<in> V\<close> t01 \<open>n \<le> N\<close> | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1473 | apply (simp add: nN_in_kkt \<open>y \<in> U\<close> inUS field_split_simps) | 
| 64792 | 1474 | done | 
| 1475 | also have "... \<subseteq> UU (X t)" | |
| 1476 | using him t01 by blast | |
| 1477 | finally show "p (k (real n / real N, y)) \<in> UU (X t)" . | |
| 1478 | qed | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1479 | with \<V> have "k (real n / real N, y) \<in> \<Union>\<V>" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1480 | by blast | 
| 64792 | 1481 | then obtain W where W: "k (real n / real N, y) \<in> W" and "W \<in> \<V>" | 
| 1482 | by blast | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1483 | then obtain p' where opeC': "openin (top_of_set C) W" | 
| 64792 | 1484 | and hom': "homeomorphism W (UU (X t)) p p'" | 
| 1485 | using homuu opeC by blast | |
| 1486 | then have "W \<subseteq> C" | |
| 1487 | using openin_imp_subset by blast | |
| 1488 | define W' where "W' = UU(X t)" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1489 | have opeVW: "openin (top_of_set V) (V \<inter> (k \<circ> Pair (n / N)) -` W)" | 
| 64792 | 1490 | apply (rule continuous_openin_preimage [OF _ _ opeC']) | 
| 1491 | apply (intro continuous_intros continuous_on_subset [OF contk]) | |
| 1492 | using kim apply (auto simp: \<open>y \<in> V\<close> W) | |
| 1493 | done | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1494 | obtain N' where opeUN': "openin (top_of_set U) N'" | 
| 64792 | 1495 |                   and "y \<in> N'" and kimw: "k ` ({(real n / real N)} \<times> N') \<subseteq> W"
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1496 | apply (rule_tac N' = "(V \<inter> (k \<circ> Pair (n / N)) -` W)" in that) | 
| 64792 | 1497 | apply (fastforce simp: \<open>y \<in> V\<close> W intro!: openin_trans [OF opeVW opeUV])+ | 
| 1498 | done | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1499 | obtain Q Q' where opeUQ: "openin (top_of_set U) Q" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1500 | and cloUQ': "closedin (top_of_set U) Q'" | 
| 64792 | 1501 | and "y \<in> Q" "Q \<subseteq> Q'" | 
| 1502 | and Q': "Q' \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V" | |
| 1503 | proof - | |
| 1504 | obtain VO VX where "open VO" "open VX" and VO: "V = U \<inter> VO" and VX: "N' = U \<inter> VX" | |
| 1505 | using opeUV opeUN' by (auto simp: openin_open) | |
| 1506 | then have "open (NN(t) \<inter> VO \<inter> VX)" | |
| 1507 | using NN t01 by blast | |
| 1508 | then obtain e where "e > 0" and e: "cball y e \<subseteq> NN(t) \<inter> VO \<inter> VX" | |
| 1509 | by (metis Int_iff \<open>N' = U \<inter> VX\<close> \<open>V = U \<inter> VO\<close> \<open>y \<in> N'\<close> \<open>y \<in> V\<close> inUS open_contains_cball t01) | |
| 1510 | show ?thesis | |
| 1511 | proof | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1512 | show "openin (top_of_set U) (U \<inter> ball y e)" | 
| 64792 | 1513 | by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1514 | show "closedin (top_of_set U) (U \<inter> cball y e)" | 
| 64792 | 1515 | using e by (auto simp: closedin_closed) | 
| 1516 | qed (use \<open>y \<in> U\<close> \<open>e > 0\<close> VO VX e in auto) | |
| 1517 | qed | |
| 1518 | then have "y \<in> Q'" "Q \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V" | |
| 1519 | by blast+ | |
| 1520 |       have neq: "{0..real n / real N} \<union> {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}"
 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1521 | apply (auto simp: field_split_simps) | 
| 64792 | 1522 | by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1) | 
| 1523 |       then have neqQ': "{0..real n / real N} \<times> Q' \<union> {real n / real N..(1 + real n) / real N} \<times> Q' = {0..(1 + real n) / real N} \<times> Q'"
 | |
| 1524 | by blast | |
| 1525 |       have cont: "continuous_on ({0..(1 + real n) / real N} \<times> Q')
 | |
| 1526 |         (\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
 | |
| 1527 | unfolding neqQ' [symmetric] | |
| 1528 | proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1529 |         show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
 | 
| 64792 | 1530 |                        ({0..real n / real N} \<times> Q')"
 | 
| 1531 | apply (simp add: closedin_closed) | |
| 1532 |           apply (rule_tac x="{0 .. real n / real N} \<times> UNIV" in exI)
 | |
| 1533 | using n_div_N_in apply (auto simp: closed_Times) | |
| 1534 | done | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1535 |         show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
 | 
| 64792 | 1536 |                        ({real n / real N..(1 + real n) / real N} \<times> Q')"
 | 
| 1537 | apply (simp add: closedin_closed) | |
| 1538 |           apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \<times> UNIV" in exI)
 | |
| 1539 | apply (auto simp: closed_Times) | |
| 1540 | by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) | |
| 1541 |         show "continuous_on ({0..real n / real N} \<times> Q') k"
 | |
| 1542 | apply (rule continuous_on_subset [OF contk]) | |
| 1543 | using Q' by auto | |
| 1544 |         have "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') h"
 | |
| 1545 | proof (rule continuous_on_subset [OF conth]) | |
| 1546 |           show "{real n / real N..(1 + real n) / real N} \<times> Q' \<subseteq> {0..1} \<times> U"
 | |
| 1547 | using \<open>N > 0\<close> | |
| 1548 | apply auto | |
| 1549 | apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) | |
| 1550 | using Suc.prems order_trans apply fastforce | |
| 1551 | apply (metis IntE cloUQ' closedin_closed) | |
| 1552 | done | |
| 1553 | qed | |
| 1554 |         moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \<times> Q')) p'"
 | |
| 1555 | proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']]) | |
| 1556 |           have "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
 | |
| 1557 | apply (rule image_mono) | |
| 1558 | using \<open>0 < \<delta>\<close> \<open>N > 0\<close> Suc.prems apply auto | |
| 1559 | apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) | |
| 1560 | using Suc.prems order_trans apply fastforce | |
| 1561 | using t Q' apply auto | |
| 1562 | done | |
| 1563 |           with him show "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> UU (X t)"
 | |
| 1564 | using t01 by blast | |
| 1565 | qed | |
| 1566 |         ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') (p' \<circ> h)"
 | |
| 1567 | by (rule continuous_on_compose) | |
| 1568 | have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \<in> Q'" for b | |
| 1569 | proof - | |
| 1570 | have "k (real n / real N, b) \<in> W" | |
| 1571 | using that Q' kimw by force | |
| 1572 | then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))" | |
| 1573 | by (simp add: homeomorphism_apply1 [OF hom']) | |
| 1574 | then show ?thesis | |
| 1575 | using Q' that by (force simp: heq) | |
| 1576 | qed | |
| 1577 |         then show "\<And>x. x \<in> {real n / real N..(1 + real n) / real N} \<times> Q' \<and>
 | |
| 1578 |                   x \<in> {0..real n / real N} \<times> Q' \<Longrightarrow> k x = (p' \<circ> h) x"
 | |
| 1579 | by auto | |
| 1580 | qed | |
| 1581 | have h_in_UU: "h (x, y) \<in> UU (X t)" if "y \<in> Q" "\<not> x \<le> real n / real N" "0 \<le> x" "x \<le> (1 + real n) / real N" for x y | |
| 1582 | proof - | |
| 1583 | have "x \<le> 1" | |
| 1584 | using Suc.prems that order_trans by force | |
| 1585 | moreover have "x \<in> K t" | |
| 1586 | by (meson atLeastAtMost_iff le_less not_le subset_eq t that) | |
| 1587 | moreover have "y \<in> U" | |
| 1588 | using \<open>y \<in> Q\<close> opeUQ openin_imp_subset by blast | |
| 1589 | moreover have "y \<in> NN t" | |
| 1590 | using Q' \<open>Q \<subseteq> Q'\<close> \<open>y \<in> Q\<close> by auto | |
| 1591 |         ultimately have "(x, y) \<in> (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
 | |
| 1592 | using that by auto | |
| 1593 |         then have "h (x, y) \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
 | |
| 1594 | by blast | |
| 1595 | also have "... \<subseteq> UU (X t)" | |
| 1596 | by (metis him t01) | |
| 1597 | finally show ?thesis . | |
| 1598 | qed | |
| 1599 |       let ?k = "(\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
 | |
| 1600 | show ?case | |
| 1601 | proof (intro exI conjI) | |
| 1602 |         show "continuous_on ({0..real (Suc n) / real N} \<times> Q) ?k"
 | |
| 1603 | apply (rule continuous_on_subset [OF cont]) | |
| 1604 | using \<open>Q \<subseteq> Q'\<close> by auto | |
| 1605 | have "\<And>a b. \<lbrakk>a \<le> real n / real N; b \<in> Q'; 0 \<le> a\<rbrakk> \<Longrightarrow> k (a, b) \<in> C" | |
| 1606 | using kim Q' by force | |
| 1607 | moreover have "\<And>a b. \<lbrakk>b \<in> Q; 0 \<le> a; a \<le> (1 + real n) / real N; \<not> a \<le> real n / real N\<rbrakk> \<Longrightarrow> p' (h (a, b)) \<in> C" | |
| 1608 | apply (rule \<open>W \<subseteq> C\<close> [THEN subsetD]) | |
| 1609 | using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \<open>Q \<subseteq> Q'\<close> \<open>W \<subseteq> C\<close> | |
| 1610 | apply auto | |
| 1611 | done | |
| 1612 |         ultimately show "?k ` ({0..real (Suc n) / real N} \<times> Q) \<subseteq> C"
 | |
| 1613 | using Q' \<open>Q \<subseteq> Q'\<close> by force | |
| 1614 | show "\<forall>z\<in>Q. ?k (0, z) = f z" | |
| 1615 | using Q' keq \<open>Q \<subseteq> Q'\<close> by auto | |
| 1616 |         show "\<forall>z \<in> {0..real (Suc n) / real N} \<times> Q. h z = p(?k z)"
 | |
| 1617 | using \<open>Q \<subseteq> U \<inter> NN t \<inter> N' \<inter> V\<close> heq apply clarsimp | |
| 1618 | using h_in_UU Q' \<open>Q \<subseteq> Q'\<close> apply (auto simp: homeomorphism_apply2 [OF hom', symmetric]) | |
| 1619 | done | |
| 1620 | qed (auto simp: \<open>y \<in> Q\<close> opeUQ) | |
| 1621 | qed | |
| 1622 | show ?thesis | |
| 1623 | using*[OF order_refl] N \<open>0 < \<delta>\<close> by (simp add: split: if_split_asm) | |
| 1624 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1625 | then obtain V fs where opeV: "\<And>y. y \<in> U \<Longrightarrow> openin (top_of_set U) (V y)" | 
| 64792 | 1626 | and V: "\<And>y. y \<in> U \<Longrightarrow> y \<in> V y" | 
| 1627 |           and contfs: "\<And>y. y \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V y) (fs y)"
 | |
| 1628 |           and *: "\<And>y. y \<in> U \<Longrightarrow> (fs y) ` ({0..1} \<times> V y) \<subseteq> C \<and>
 | |
| 1629 | (\<forall>z \<in> V y. fs y (0, z) = f z) \<and> | |
| 1630 |                             (\<forall>z \<in> {0..1} \<times> V y. h z = p(fs y z))"
 | |
| 1631 | by (metis (mono_tags)) | |
| 1632 | then have VU: "\<And>y. y \<in> U \<Longrightarrow> V y \<subseteq> U" | |
| 1633 | by (meson openin_imp_subset) | |
| 1634 |   obtain k where contk: "continuous_on ({0..1} \<times> U) k"
 | |
| 1635 |              and k: "\<And>x i. \<lbrakk>i \<in> U; x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i\<rbrakk> \<Longrightarrow> k x = fs i x"
 | |
| 1636 | proof (rule pasting_lemma_exists) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1637 |     let ?X = "top_of_set ({0..1::real} \<times> U)"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1638 |     show "topspace ?X \<subseteq> (\<Union>i\<in>U. {0..1} \<times> V i)"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1639 | using V by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1640 |     show "\<And>i. i \<in> U \<Longrightarrow> openin (top_of_set ({0..1} \<times> U)) ({0..1} \<times> V i)"
 | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1641 | by (simp add: Abstract_Topology.openin_Times opeV) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1642 | show "\<And>i. i \<in> U \<Longrightarrow> continuous_map | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1643 |               (subtopology (top_of_set ({0..1} \<times> U)) ({0..1} \<times> V i)) euclidean (fs i)"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1644 | using contfs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1645 | apply simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1646 | by (metis continuous_map_iff_continuous continuous_on_subset openin_imp_subset openin_subtopology_self subtopology_subtopology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1647 |     show "fs i x = fs j x"  if "i \<in> U" "j \<in> U" and x: "x \<in> topspace ?X \<inter> {0..1} \<times> V i \<inter> {0..1} \<times> V j"
 | 
| 64792 | 1648 | for i j x | 
| 1649 | proof - | |
| 1650 | obtain u y where "x = (u, y)" "y \<in> V i" "y \<in> V j" "0 \<le> u" "u \<le> 1" | |
| 1651 | using x by auto | |
| 1652 | show ?thesis | |
| 1653 |       proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \<times> {y}" h])
 | |
| 1654 | show "fs i (0, y) = fs j (0, y)" | |
| 1655 | using*V by (simp add: \<open>y \<in> V i\<close> \<open>y \<in> V j\<close> that) | |
| 1656 |         show conth_y: "continuous_on ({0..1} \<times> {y}) h"
 | |
| 1657 | apply (rule continuous_on_subset [OF conth]) | |
| 1658 | using VU \<open>y \<in> V j\<close> that by auto | |
| 1659 |         show "h ` ({0..1} \<times> {y}) \<subseteq> S"
 | |
| 1660 | using \<open>y \<in> V i\<close> assms(3) VU that by fastforce | |
| 1661 |         show "continuous_on ({0..1} \<times> {y}) (fs i)"
 | |
| 1662 | using continuous_on_subset [OF contfs] \<open>i \<in> U\<close> | |
| 1663 | by (simp add: \<open>y \<in> V i\<close> subset_iff) | |
| 1664 |         show "fs i ` ({0..1} \<times> {y}) \<subseteq> C"
 | |
| 1665 | using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by fastforce | |
| 1666 |         show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs i x)"
 | |
| 1667 | using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by blast | |
| 1668 |         show "continuous_on ({0..1} \<times> {y}) (fs j)"
 | |
| 1669 | using continuous_on_subset [OF contfs] \<open>j \<in> U\<close> | |
| 1670 | by (simp add: \<open>y \<in> V j\<close> subset_iff) | |
| 1671 |         show "fs j ` ({0..1} \<times> {y}) \<subseteq> C"
 | |
| 1672 | using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by fastforce | |
| 1673 |         show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs j x)"
 | |
| 1674 | using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by blast | |
| 1675 |         show "connected ({0..1::real} \<times> {y})"
 | |
| 1676 | using connected_Icc connected_Times connected_sing by blast | |
| 1677 |         show "(0, y) \<in> {0..1::real} \<times> {y}"
 | |
| 1678 | by force | |
| 1679 |         show "x \<in> {0..1} \<times> {y}"
 | |
| 1680 | using \<open>x = (u, y)\<close> x by blast | |
| 1681 | qed | |
| 1682 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 1683 | qed force | 
| 64792 | 1684 | show ?thesis | 
| 1685 | proof | |
| 1686 |     show "k ` ({0..1} \<times> U) \<subseteq> C"
 | |
| 1687 | using V*k VU by fastforce | |
| 1688 | show "\<And>y. y \<in> U \<Longrightarrow> k (0, y) = f y" | |
| 1689 | by (simp add: V*k) | |
| 1690 |     show "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p (k z)"
 | |
| 1691 | using V*k by auto | |
| 1692 | qed (auto simp: contk) | |
| 1693 | qed | |
| 1694 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1695 | corollary covering_space_lift_homotopy_alt: | 
| 64792 | 1696 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1697 | and h :: "'c::real_normed_vector \<times> real \<Rightarrow> 'b" | |
| 1698 | assumes cov: "covering_space C p S" | |
| 1699 |       and conth: "continuous_on (U \<times> {0..1}) h"
 | |
| 1700 |       and him: "h ` (U \<times> {0..1}) \<subseteq> S"
 | |
| 1701 | and heq: "\<And>y. y \<in> U \<Longrightarrow> h (y,0) = p(f y)" | |
| 1702 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C" | |
| 1703 |   obtains k where "continuous_on (U \<times> {0..1}) k"
 | |
| 1704 |                   "k ` (U \<times> {0..1}) \<subseteq> C"
 | |
| 1705 | "\<And>y. y \<in> U \<Longrightarrow> k(y, 0) = f y" | |
| 1706 |                   "\<And>z. z \<in> U \<times> {0..1} \<Longrightarrow> h z = p(k z)"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1707 | proof - | 
| 64792 | 1708 |   have "continuous_on ({0..1} \<times> U) (h \<circ> (\<lambda>z. (snd z, fst z)))"
 | 
| 1709 | by (intro continuous_intros continuous_on_subset [OF conth]) auto | |
| 1710 |   then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
 | |
| 1711 |                   and kim:  "k ` ({0..1} \<times> U) \<subseteq> C"
 | |
| 1712 | and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y" | |
| 1713 |                   and heqp: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> (h \<circ> (\<lambda>z. Pair (snd z) (fst z))) z = p(k z)"
 | |
| 1714 | apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim]) | |
| 1715 | using him by (auto simp: contf heq) | |
| 1716 | show ?thesis | |
| 1717 | apply (rule_tac k="k \<circ> (\<lambda>z. Pair (snd z) (fst z))" in that) | |
| 1718 | apply (intro continuous_intros continuous_on_subset [OF contk]) | |
| 1719 | using kim heqp apply (auto simp: k0) | |
| 1720 | done | |
| 1721 | qed | |
| 1722 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1723 | corollary covering_space_lift_homotopic_function: | 
| 64792 | 1724 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and g:: "'c::real_normed_vector \<Rightarrow> 'a" | 
| 1725 | assumes cov: "covering_space C p S" | |
| 1726 | and contg: "continuous_on U g" | |
| 1727 | and gim: "g ` U \<subseteq> C" | |
| 1728 | and pgeq: "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1729 | and hom: "homotopic_with_canon (\<lambda>x. True) U S f f'" | 
| 64792 | 1730 | obtains g' where "continuous_on U g'" "image g' U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g' y) = f' y" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1731 | proof - | 
| 64792 | 1732 |   obtain h where conth: "continuous_on ({0..1::real} \<times> U) h"
 | 
| 1733 |              and him: "h ` ({0..1} \<times> U) \<subseteq> S"
 | |
| 1734 | and h0: "\<And>x. h(0, x) = f x" | |
| 1735 | and h1: "\<And>x. h(1, x) = f' x" | |
| 1736 | using hom by (auto simp: homotopic_with_def) | |
| 1737 | have "\<And>y. y \<in> U \<Longrightarrow> h (0, y) = p (g y)" | |
| 1738 | by (simp add: h0 pgeq) | |
| 1739 |   then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
 | |
| 1740 |                   and kim: "k ` ({0..1} \<times> U) \<subseteq> C"
 | |
| 1741 | and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = g y" | |
| 1742 |                   and heq: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
 | |
| 1743 | using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis | |
| 1744 | show ?thesis | |
| 1745 | proof | |
| 1746 | show "continuous_on U (k \<circ> Pair 1)" | |
| 1747 | by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one) | |
| 1748 | show "(k \<circ> Pair 1) ` U \<subseteq> C" | |
| 1749 | using kim by auto | |
| 1750 | show "\<And>y. y \<in> U \<Longrightarrow> p ((k \<circ> Pair 1) y) = f' y" | |
| 1751 | by (auto simp: h1 heq [symmetric]) | |
| 1752 | qed | |
| 1753 | qed | |
| 1754 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1755 | corollary covering_space_lift_inessential_function: | 
| 64792 | 1756 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and U :: "'c::real_normed_vector set" | 
| 1757 | assumes cov: "covering_space C p S" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1758 | and hom: "homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. a)" | 
| 64792 | 1759 | obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1760 | proof (cases "U = {}")
 | 
| 64792 | 1761 | case True | 
| 1762 | then show ?thesis | |
| 1763 | using that continuous_on_empty by blast | |
| 1764 | next | |
| 1765 | case False | |
| 1766 | then obtain b where b: "b \<in> C" "p b = a" | |
| 1767 | using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom] | |
| 1768 | by auto | |
| 1769 | then have gim: "(\<lambda>y. b) ` U \<subseteq> C" | |
| 1770 | by blast | |
| 1771 | show ?thesis | |
| 1772 | apply (rule covering_space_lift_homotopic_function | |
| 1773 | [OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]]) | |
| 1774 | using b that apply auto | |
| 1775 | done | |
| 1776 | qed | |
| 1777 | ||
| 69683 | 1778 | subsection\<open> Lifting of general functions to covering space\<close> | 
| 64792 | 1779 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1780 | proposition covering_space_lift_path_strong: | 
| 64792 | 1781 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1782 | and f :: "'c::real_normed_vector \<Rightarrow> 'b" | |
| 1783 | assumes cov: "covering_space C p S" and "a \<in> C" | |
| 1784 | and "path g" and pag: "path_image g \<subseteq> S" and pas: "pathstart g = p a" | |
| 1785 | obtains h where "path h" "path_image h \<subseteq> C" "pathstart h = a" | |
| 1786 |                 and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1787 | proof - | 
| 64792 | 1788 | obtain k:: "real \<times> 'c \<Rightarrow> 'a" | 
| 1789 |     where contk: "continuous_on ({0..1} \<times> {undefined}) k"
 | |
| 1790 |       and kim: "k ` ({0..1} \<times> {undefined}) \<subseteq> C"
 | |
| 1791 | and k0: "k (0, undefined) = a" | |
| 1792 |       and pk: "\<And>z. z \<in> {0..1} \<times> {undefined} \<Longrightarrow> p(k z) = (g \<circ> fst) z"
 | |
| 1793 |   proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \<circ> fst"])
 | |
| 1794 |     show "continuous_on ({0..1::real} \<times> {undefined::'c}) (g \<circ> fst)"
 | |
| 1795 | apply (intro continuous_intros) | |
| 1796 | using \<open>path g\<close> by (simp add: path_def) | |
| 1797 |     show "(g \<circ> fst) ` ({0..1} \<times> {undefined}) \<subseteq> S"
 | |
| 1798 | using pag by (auto simp: path_image_def) | |
| 1799 |     show "(g \<circ> fst) (0, y) = p a" if "y \<in> {undefined}" for y::'c
 | |
| 1800 | by (metis comp_def fst_conv pas pathstart_def) | |
| 1801 | qed (use assms in auto) | |
| 1802 | show ?thesis | |
| 1803 | proof | |
| 1804 | show "path (k \<circ> (\<lambda>t. Pair t undefined))" | |
| 1805 | unfolding path_def | |
| 1806 | by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto | |
| 1807 | show "path_image (k \<circ> (\<lambda>t. (t, undefined))) \<subseteq> C" | |
| 1808 | using kim by (auto simp: path_image_def) | |
| 1809 | show "pathstart (k \<circ> (\<lambda>t. (t, undefined))) = a" | |
| 1810 | by (auto simp: pathstart_def k0) | |
| 1811 |     show "\<And>t. t \<in> {0..1} \<Longrightarrow> p ((k \<circ> (\<lambda>t. (t, undefined))) t) = g t"
 | |
| 1812 | by (auto simp: pk) | |
| 1813 | qed | |
| 1814 | qed | |
| 1815 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1816 | corollary covering_space_lift_path: | 
| 64792 | 1817 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1818 | assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \<subseteq> S" | |
| 1819 |   obtains h where "path h" "path_image h \<subseteq> C" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1820 | proof - | 
| 64792 | 1821 | obtain a where "a \<in> C" "pathstart g = p a" | 
| 1822 | by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE) | |
| 1823 | show ?thesis | |
| 1824 | using covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close> \<open>path g\<close> pig] | |
| 1825 | by (metis \<open>pathstart g = p a\<close> that) | |
| 1826 | qed | |
| 1827 | ||
| 1828 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1829 | proposition covering_space_lift_homotopic_paths: | 
| 64792 | 1830 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1831 | assumes cov: "covering_space C p S" | |
| 1832 | and "path g1" and pig1: "path_image g1 \<subseteq> S" | |
| 1833 | and "path g2" and pig2: "path_image g2 \<subseteq> S" | |
| 1834 | and hom: "homotopic_paths S g1 g2" | |
| 1835 |       and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
 | |
| 1836 |       and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
 | |
| 1837 | and h1h2: "pathstart h1 = pathstart h2" | |
| 1838 | shows "homotopic_paths C h1 h2" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1839 | proof - | 
| 64792 | 1840 | obtain h :: "real \<times> real \<Rightarrow> 'b" | 
| 1841 |      where conth: "continuous_on ({0..1} \<times> {0..1}) h"
 | |
| 1842 |        and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
 | |
| 1843 | and h0: "\<And>x. h (0, x) = g1 x" and h1: "\<And>x. h (1, x) = g2 x" | |
| 1844 |        and heq0: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 0) = g1 0"
 | |
| 1845 |        and heq1: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 1) = g1 1"
 | |
| 1846 | using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def) | |
| 1847 |   obtain k where contk: "continuous_on ({0..1} \<times> {0..1}) k"
 | |
| 1848 |              and kim: "k ` ({0..1} \<times> {0..1}) \<subseteq> C"
 | |
| 1849 |              and kh2: "\<And>y. y \<in> {0..1} \<Longrightarrow> k (y, 0) = h2 0"
 | |
| 1850 |              and hpk: "\<And>z. z \<in> {0..1} \<times> {0..1} \<Longrightarrow> h z = p (k z)"
 | |
| 1851 | apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\<lambda>x. h2 0"]) | |
| 1852 | using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def) | |
| 71172 | 1853 | using path_image_def pih2 by fastforce+ | 
| 64792 | 1854 |   have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2"
 | 
| 1855 | using \<open>path g1\<close> \<open>path g2\<close> path_def by blast+ | |
| 1856 |   have g1im: "g1 ` {0..1} \<subseteq> S" and g2im: "g2 ` {0..1} \<subseteq> S"
 | |
| 1857 | using path_image_def pig1 pig2 by auto | |
| 1858 |   have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2"
 | |
| 1859 | using \<open>path h1\<close> \<open>path h2\<close> path_def by blast+ | |
| 1860 |   have h1im: "h1 ` {0..1} \<subseteq> C" and h2im: "h2 ` {0..1} \<subseteq> C"
 | |
| 1861 | using path_image_def pih1 pih2 by auto | |
| 1862 | show ?thesis | |
| 1863 | unfolding homotopic_paths pathstart_def pathfinish_def | |
| 1864 | proof (intro exI conjI ballI) | |
| 1865 |     show keqh1: "k(0, x) = h1 x" if "x \<in> {0..1}" for x
 | |
| 1866 | proof (rule covering_space_lift_unique [OF cov _ contg1 g1im]) | |
| 1867 | show "k (0,0) = h1 0" | |
| 1868 | by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one) | |
| 1869 |       show "continuous_on {0..1} (\<lambda>a. k (0, a))"
 | |
| 1870 | by (intro continuous_intros continuous_on_compose2 [OF contk]) auto | |
| 1871 |       show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 x = p (k (0, x))"
 | |
| 1872 | by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl) | |
| 1873 | qed (use conth1 h1im kim that in \<open>auto simp: ph1\<close>) | |
| 1874 |     show "k(1, x) = h2 x" if "x \<in> {0..1}" for x
 | |
| 1875 | proof (rule covering_space_lift_unique [OF cov _ contg2 g2im]) | |
| 1876 | show "k (1,0) = h2 0" | |
| 1877 | by (metis atLeastAtMost_iff kh2 order_refl zero_le_one) | |
| 1878 |       show "continuous_on {0..1} (\<lambda>a. k (1, a))"
 | |
| 1879 | by (intro continuous_intros continuous_on_compose2 [OF contk]) auto | |
| 1880 |       show "\<And>x. x \<in> {0..1} \<Longrightarrow> g2 x = p (k (1, x))"
 | |
| 1881 | by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one) | |
| 1882 | qed (use conth2 h2im kim that in \<open>auto simp: ph2\<close>) | |
| 1883 |     show "\<And>t. t \<in> {0..1} \<Longrightarrow> (k \<circ> Pair t) 0 = h1 0"
 | |
| 1884 | by (metis comp_apply h1h2 kh2 pathstart_def) | |
| 1885 |     show "(k \<circ> Pair t) 1 = h1 1" if "t \<in> {0..1}" for t
 | |
| 1886 | proof (rule covering_space_lift_unique | |
| 1887 |            [OF cov, of "\<lambda>a. (k \<circ> Pair a) 1" 0 "\<lambda>a. h1 1" "{0..1}"  "\<lambda>x. g1 1"])
 | |
| 1888 | show "(k \<circ> Pair 0) 1 = h1 1" | |
| 1889 | using keqh1 by auto | |
| 1890 |       show "continuous_on {0..1} (\<lambda>a. (k \<circ> Pair a) 1)"
 | |
| 1891 | apply simp | |
| 1892 | by (intro continuous_intros continuous_on_compose2 [OF contk]) auto | |
| 1893 |       show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 1 = p ((k \<circ> Pair x) 1)"
 | |
| 1894 | using heq1 hpk by auto | |
| 71172 | 1895 | qed (use contk kim g1im h1im that in \<open>auto simp: ph1\<close>) | 
| 64792 | 1896 | qed (use contk kim in auto) | 
| 1897 | qed | |
| 1898 | ||
| 1899 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1900 | corollary covering_space_monodromy: | 
| 64792 | 1901 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1902 | assumes cov: "covering_space C p S" | |
| 1903 | and "path g1" and pig1: "path_image g1 \<subseteq> S" | |
| 1904 | and "path g2" and pig2: "path_image g2 \<subseteq> S" | |
| 1905 | and hom: "homotopic_paths S g1 g2" | |
| 1906 |       and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
 | |
| 1907 |       and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
 | |
| 1908 | and h1h2: "pathstart h1 = pathstart h2" | |
| 1909 | shows "pathfinish h1 = pathfinish h2" | |
| 70136 | 1910 | using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish | 
| 1911 | by blast | |
| 64792 | 1912 | |
| 1913 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1914 | corollary covering_space_lift_homotopic_path: | 
| 64792 | 1915 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1916 | assumes cov: "covering_space C p S" | |
| 1917 | and hom: "homotopic_paths S f f'" | |
| 1918 | and "path g" and pig: "path_image g \<subseteq> C" | |
| 1919 | and a: "pathstart g = a" and b: "pathfinish g = b" | |
| 1920 |       and pgeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g t) = f t"
 | |
| 1921 | obtains g' where "path g'" "path_image g' \<subseteq> C" | |
| 1922 |                    "pathstart g' = a" "pathfinish g' = b" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g' t) = f' t"
 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1923 | proof (rule covering_space_lift_path_strong [OF cov, of a f']) | 
| 64792 | 1924 | show "a \<in> C" | 
| 1925 | using a pig by auto | |
| 1926 | show "path f'" "path_image f' \<subseteq> S" | |
| 1927 | using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+ | |
| 1928 | show "pathstart f' = p a" | |
| 1929 | by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one) | |
| 1930 | qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig) | |
| 1931 | ||
| 1932 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1933 | proposition covering_space_lift_general: | 
| 64792 | 1934 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 1935 | and f :: "'c::real_normed_vector \<Rightarrow> 'b" | |
| 1936 | assumes cov: "covering_space C p S" and "a \<in> C" "z \<in> U" | |
| 1937 | and U: "path_connected U" "locally path_connected U" | |
| 1938 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S" | |
| 1939 | and feq: "f z = p a" | |
| 1940 | and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk> | |
| 1941 | \<Longrightarrow> \<exists>q. path q \<and> path_image q \<subseteq> C \<and> | |
| 1942 | pathstart q = a \<and> pathfinish q = a \<and> | |
| 1943 | homotopic_paths S (f \<circ> r) (p \<circ> q)" | |
| 1944 | obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1945 | proof - | 
| 64792 | 1946 | have *: "\<exists>g h. path g \<and> path_image g \<subseteq> U \<and> | 
| 1947 | pathstart g = z \<and> pathfinish g = y \<and> | |
| 1948 | path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and> | |
| 1949 |                  (\<forall>t \<in> {0..1}. p(h t) = f(g t))"
 | |
| 1950 | if "y \<in> U" for y | |
| 1951 | proof - | |
| 1952 | obtain g where "path g" "path_image g \<subseteq> U" and pastg: "pathstart g = z" | |
| 1953 | and pafig: "pathfinish g = y" | |
| 1954 | using U \<open>z \<in> U\<close> \<open>y \<in> U\<close> by (force simp: path_connected_def) | |
| 1955 | obtain h where "path h" "path_image h \<subseteq> C" "pathstart h = a" | |
| 1956 |                and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = (f \<circ> g) t"
 | |
| 1957 | proof (rule covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close>]) | |
| 1958 | show "path (f \<circ> g)" | |
| 1959 | using \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> contf continuous_on_subset path_continuous_image by blast | |
| 1960 | show "path_image (f \<circ> g) \<subseteq> S" | |
| 1961 | by (metis \<open>path_image g \<subseteq> U\<close> fim image_mono path_image_compose subset_trans) | |
| 1962 | show "pathstart (f \<circ> g) = p a" | |
| 1963 | by (simp add: feq pastg pathstart_compose) | |
| 1964 | qed auto | |
| 1965 | then show ?thesis | |
| 1966 | by (metis \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> comp_apply pafig pastg) | |
| 1967 | qed | |
| 1968 | have "\<exists>l. \<forall>g h. path g \<and> path_image g \<subseteq> U \<and> pathstart g = z \<and> pathfinish g = y \<and> | |
| 1969 | path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and> | |
| 1970 |                   (\<forall>t \<in> {0..1}. p(h t) = f(g t))  \<longrightarrow> pathfinish h = l" for y
 | |
| 1971 | proof - | |
| 1972 | have "pathfinish h = pathfinish h'" | |
| 1973 | if g: "path g" "path_image g \<subseteq> U" "pathstart g = z" "pathfinish g = y" | |
| 1974 | and h: "path h" "path_image h \<subseteq> C" "pathstart h = a" | |
| 1975 |             and phg: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)"
 | |
| 1976 | and g': "path g'" "path_image g' \<subseteq> U" "pathstart g' = z" "pathfinish g' = y" | |
| 1977 | and h': "path h'" "path_image h' \<subseteq> C" "pathstart h' = a" | |
| 1978 |             and phg': "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h' t) = f(g' t)"
 | |
| 1979 | for g h g' h' | |
| 1980 | proof - | |
| 1981 | obtain q where "path q" and piq: "path_image q \<subseteq> C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a" | |
| 1982 | and homS: "homotopic_paths S (f \<circ> g +++ reversepath g') (p \<circ> q)" | |
| 1983 | using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join) | |
| 1984 | have papq: "path (p \<circ> q)" | |
| 1985 | using homS homotopic_paths_imp_path by blast | |
| 1986 | have pipq: "path_image (p \<circ> q) \<subseteq> S" | |
| 1987 | using homS homotopic_paths_imp_subset by blast | |
| 1988 | obtain q' where "path q'" "path_image q' \<subseteq> C" | |
| 1989 | and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q" | |
| 1990 |                 and pq'_eq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p (q' t) = (f \<circ> g +++ reversepath g') t"
 | |
| 1991 | using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \<open>path q\<close> piq refl refl] | |
| 1992 | by auto | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1993 | have "q' t = (h \<circ> (*\<^sub>R) 2) t" if "0 \<le> t" "t \<le> 1/2" for t | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1994 |       proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \<circ> (*\<^sub>R) 2" "{0..1/2}" "f \<circ> g \<circ> (*\<^sub>R) 2" t])
 | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1995 | show "q' 0 = (h \<circ> (*\<^sub>R) 2) 0" | 
| 71172 | 1996 | by (metis \<open>pathstart q' = pathstart q\<close> comp_def h pastq pathstart_def pth_4(2)) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 1997 |         show "continuous_on {0..1/2} (f \<circ> g \<circ> (*\<^sub>R) 2)"
 | 
| 64792 | 1998 | apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path g\<close>] continuous_on_subset [OF contf]) | 
| 1999 | using g(2) path_image_def by fastforce+ | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 2000 |         show "(f \<circ> g \<circ> (*\<^sub>R) 2) ` {0..1/2} \<subseteq> S"
 | 
| 64792 | 2001 | using g(2) path_image_def fim by fastforce | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 2002 |         show "(h \<circ> (*\<^sub>R) 2) ` {0..1/2} \<subseteq> C"
 | 
| 64792 | 2003 | using h path_image_def by fastforce | 
| 2004 |         show "q' ` {0..1/2} \<subseteq> C"
 | |
| 2005 | using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 2006 |         show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> (*\<^sub>R) 2) x = p (q' x)"
 | 
| 64792 | 2007 | by (auto simp: joinpaths_def pq'_eq) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 2008 |         show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> (*\<^sub>R) 2) x = p ((h \<circ> (*\<^sub>R) 2) x)"
 | 
| 64792 | 2009 | by (simp add: phg) | 
| 2010 |         show "continuous_on {0..1/2} q'"
 | |
| 2011 | by (simp add: continuous_on_path \<open>path q'\<close>) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 2012 |         show "continuous_on {0..1/2} (h \<circ> (*\<^sub>R) 2)"
 | 
| 64792 | 2013 | apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path h\<close>], force) | 
| 2014 | done | |
| 2015 | qed (use that in auto) | |
| 2016 | moreover have "q' t = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \<le> 1" for t | |
| 2017 |       proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" t])
 | |
| 2018 | show "q' 1 = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) 1" | |
| 2019 | using h' \<open>pathfinish q' = pathfinish q\<close> pafiq | |
| 2020 | by (simp add: pathstart_def pathfinish_def reversepath_def) | |
| 2021 |         show "continuous_on {1/2<..1} (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
 | |
| 2022 | apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path g'\<close> continuous_on_subset [OF contf]) | |
| 2023 | using g' apply simp_all | |
| 2024 | by (auto simp: path_image_def reversepath_def) | |
| 2025 |         show "(f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> S"
 | |
| 2026 | using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def) | |
| 2027 |         show "q' ` {1/2<..1} \<subseteq> C"
 | |
| 2028 | using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce | |
| 2029 |         show "(reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> C"
 | |
| 2030 | using h' by (simp add: path_image_def reversepath_def subset_eq) | |
| 2031 |         show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow> (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p (q' x)"
 | |
| 2032 | by (auto simp: joinpaths_def pq'_eq) | |
| 2033 |         show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow>
 | |
| 2034 | (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x)" | |
| 2035 | by (simp add: phg' reversepath_def) | |
| 2036 |         show "continuous_on {1/2<..1} q'"
 | |
| 2037 | by (auto intro: continuous_on_path [OF \<open>path q'\<close>]) | |
| 2038 |         show "continuous_on {1/2<..1} (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
 | |
| 2039 | apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path h'\<close>) | |
| 2040 | using h' apply auto | |
| 2041 | done | |
| 2042 | qed (use that in auto) | |
| 2043 | ultimately have "q' t = (h +++ reversepath h') t" if "0 \<le> t" "t \<le> 1" for t | |
| 2044 | using that by (simp add: joinpaths_def) | |
| 2045 | then have "path(h +++ reversepath h')" | |
| 2046 | by (auto intro: path_eq [OF \<open>path q'\<close>]) | |
| 2047 | then show ?thesis | |
| 2048 | by (auto simp: \<open>path h\<close> \<open>path h'\<close>) | |
| 2049 | qed | |
| 2050 | then show ?thesis by metis | |
| 2051 | qed | |
| 2052 | then obtain l :: "'c \<Rightarrow> 'a" | |
| 2053 | where l: "\<And>y g h. \<lbrakk>path g; path_image g \<subseteq> U; pathstart g = z; pathfinish g = y; | |
| 2054 | path h; path_image h \<subseteq> C; pathstart h = a; | |
| 2055 |                              \<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)\<rbrakk> \<Longrightarrow> pathfinish h = l y"
 | |
| 2056 | by metis | |
| 2057 | show ?thesis | |
| 2058 | proof | |
| 2059 | show pleq: "p (l y) = f y" if "y \<in> U" for y | |
| 2060 | using*[OF \<open>y \<in> U\<close>] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one) | |
| 2061 | show "l z = a" | |
| 2062 | using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) | |
| 2063 | show LC: "l ` U \<subseteq> C" | |
| 2064 | by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2065 | have "\<exists>T. openin (top_of_set U) T \<and> y \<in> T \<and> T \<subseteq> U \<inter> l -` X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2066 | if X: "openin (top_of_set C) X" and "y \<in> U" "l y \<in> X" for X y | 
| 64792 | 2067 | proof - | 
| 2068 | have "X \<subseteq> C" | |
| 2069 | using X openin_euclidean_subtopology_iff by blast | |
| 2070 | have "f y \<in> S" | |
| 2071 | using fim \<open>y \<in> U\<close> by blast | |
| 2072 | then obtain W \<V> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2073 | where WV: "f y \<in> W \<and> openin (top_of_set S) W \<and> | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2074 | (\<Union>\<V> = C \<inter> p -` W \<and> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2075 | (\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and> | 
| 64792 | 2076 | pairwise disjnt \<V> \<and> | 
| 2077 | (\<forall>U \<in> \<V>. \<exists>q. homeomorphism U W p q))" | |
| 2078 | using cov by (force simp: covering_space_def) | |
| 2079 | then have "l y \<in> \<Union>\<V>" | |
| 2080 | using \<open>X \<subseteq> C\<close> pleq that by auto | |
| 2081 | then obtain W' where "l y \<in> W'" and "W' \<in> \<V>" | |
| 2082 | by blast | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2083 | with WV obtain p' where opeCW': "openin (top_of_set C) W'" | 
| 64792 | 2084 | and homUW': "homeomorphism W' W p p'" | 
| 2085 | by blast | |
| 2086 | then have contp': "continuous_on W p'" and p'im: "p' ` W \<subseteq> W'" | |
| 2087 | using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ | |
| 2088 | obtain V where "y \<in> V" "y \<in> U" and fimW: "f ` V \<subseteq> W" "V \<subseteq> U" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2089 | and "path_connected V" and opeUV: "openin (top_of_set U) V" | 
| 64792 | 2090 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2091 | have "openin (top_of_set U) (U \<inter> f -` W)" | 
| 64792 | 2092 | using WV contf continuous_on_open_gen fim by auto | 
| 2093 | then show ?thesis | |
| 2094 | using U WV | |
| 2095 | apply (auto simp: locally_path_connected) | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2096 | apply (drule_tac x="U \<inter> f -` W" in spec) | 
| 64792 | 2097 | apply (drule_tac x=y in spec) | 
| 2098 | apply (auto simp: \<open>y \<in> U\<close> intro: that) | |
| 2099 | done | |
| 2100 | qed | |
| 2101 | have "W' \<subseteq> C" "W \<subseteq> S" | |
| 2102 | using opeCW' WV openin_imp_subset by auto | |
| 2103 | have p'im: "p' ` W \<subseteq> W'" | |
| 2104 | using homUW' homeomorphism_image2 by fastforce | |
| 2105 | show ?thesis | |
| 2106 | proof (intro exI conjI) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2107 | have "openin (top_of_set S) (W \<inter> p' -` (W' \<inter> X))" | 
| 64792 | 2108 | proof (rule openin_trans) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2109 | show "openin (top_of_set W) (W \<inter> p' -` (W' \<inter> X))" | 
| 64792 | 2110 | apply (rule continuous_openin_preimage [OF contp' p'im]) | 
| 2111 | using X \<open>W' \<subseteq> C\<close> apply (auto simp: openin_open) | |
| 2112 | done | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2113 | show "openin (top_of_set S) W" | 
| 64792 | 2114 | using WV by blast | 
| 2115 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 2116 | then show "openin (top_of_set U) (V \<inter> (U \<inter> (f -` (W \<inter> (p' -` (W' \<inter> X))))))" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2117 | by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2118 | have "p' (f y) \<in> X" | 
| 64792 | 2119 | using \<open>l y \<in> W'\<close> homeomorphism_apply1 [OF homUW'] pleq \<open>y \<in> U\<close> \<open>l y \<in> X\<close> by fastforce | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2120 | then show "y \<in> V \<inter> (U \<inter> f -` (W \<inter> p' -` (W' \<inter> X)))" | 
| 64792 | 2121 | using \<open>y \<in> U\<close> \<open>y \<in> V\<close> WV p'im by auto | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2122 | show "V \<inter> (U \<inter> f -` (W \<inter> p' -` (W' \<inter> X))) \<subseteq> U \<inter> l -` X" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2123 | proof (intro subsetI IntI; clarify) | 
| 64792 | 2124 | fix y' | 
| 2125 | assume y': "y' \<in> V" "y' \<in> U" "f y' \<in> W" "p' (f y') \<in> W'" "p' (f y') \<in> X" | |
| 2126 | then obtain \<gamma> where "path \<gamma>" "path_image \<gamma> \<subseteq> V" "pathstart \<gamma> = y" "pathfinish \<gamma> = y'" | |
| 2127 | by (meson \<open>path_connected V\<close> \<open>y \<in> V\<close> path_connected_def) | |
| 2128 | obtain pp qq where "path pp" "path_image pp \<subseteq> U" | |
| 2129 | "pathstart pp = z" "pathfinish pp = y" | |
| 2130 | "path qq" "path_image qq \<subseteq> C" "pathstart qq = a" | |
| 2131 |                          and pqqeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(qq t) = f(pp t)"
 | |
| 2132 | using*[OF \<open>y \<in> U\<close>] by blast | |
| 2133 | have finW: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> f (\<gamma> x) \<in> W" | |
| 2134 | using \<open>path_image \<gamma> \<subseteq> V\<close> by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD]) | |
| 2135 | have "pathfinish (qq +++ (p' \<circ> f \<circ> \<gamma>)) = l y'" | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2136 | proof (rule l [of "pp +++ \<gamma>" y' "qq +++ (p' \<circ> f \<circ> \<gamma>)"]) | 
| 64792 | 2137 | show "path (pp +++ \<gamma>)" | 
| 2138 | by (simp add: \<open>path \<gamma>\<close> \<open>path pp\<close> \<open>pathfinish pp = y\<close> \<open>pathstart \<gamma> = y\<close>) | |
| 2139 | show "path_image (pp +++ \<gamma>) \<subseteq> U" | |
| 2140 | using \<open>V \<subseteq> U\<close> \<open>path_image \<gamma> \<subseteq> V\<close> \<open>path_image pp \<subseteq> U\<close> not_in_path_image_join by blast | |
| 2141 | show "pathstart (pp +++ \<gamma>) = z" | |
| 2142 | by (simp add: \<open>pathstart pp = z\<close>) | |
| 2143 | show "pathfinish (pp +++ \<gamma>) = y'" | |
| 2144 | by (simp add: \<open>pathfinish \<gamma> = y'\<close>) | |
| 2145 | have paqq: "pathfinish qq = pathstart (p' \<circ> f \<circ> \<gamma>)" | |
| 2146 | apply (simp add: \<open>pathstart \<gamma> = y\<close> pathstart_compose) | |
| 2147 | apply (metis (mono_tags, lifting) \<open>l y \<in> W'\<close> \<open>path pp\<close> \<open>path qq\<close> \<open>path_image pp \<subseteq> U\<close> \<open>path_image qq \<subseteq> C\<close> | |
| 2148 | \<open>pathfinish pp = y\<close> \<open>pathstart pp = z\<close> \<open>pathstart qq = a\<close> | |
| 2149 | homeomorphism_apply1 [OF homUW'] l pleq pqqeq \<open>y \<in> U\<close>) | |
| 2150 | done | |
| 2151 | have "continuous_on (path_image \<gamma>) (p' \<circ> f)" | |
| 2152 | proof (rule continuous_on_compose) | |
| 2153 | show "continuous_on (path_image \<gamma>) f" | |
| 2154 | using \<open>path_image \<gamma> \<subseteq> V\<close> \<open>V \<subseteq> U\<close> contf continuous_on_subset by blast | |
| 2155 | show "continuous_on (f ` path_image \<gamma>) p'" | |
| 2156 | apply (rule continuous_on_subset [OF contp']) | |
| 2157 | apply (auto simp: path_image_def pathfinish_def pathstart_def finW) | |
| 2158 | done | |
| 2159 | qed | |
| 2160 | then show "path (qq +++ (p' \<circ> f \<circ> \<gamma>))" | |
| 2161 | using \<open>path \<gamma>\<close> \<open>path qq\<close> paqq path_continuous_image path_join_imp by blast | |
| 2162 | show "path_image (qq +++ (p' \<circ> f \<circ> \<gamma>)) \<subseteq> C" | |
| 2163 | apply (rule subset_path_image_join) | |
| 2164 | apply (simp add: \<open>path_image qq \<subseteq> C\<close>) | |
| 2165 | by (metis \<open>W' \<subseteq> C\<close> \<open>path_image \<gamma> \<subseteq> V\<close> dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose) | |
| 2166 | show "pathstart (qq +++ (p' \<circ> f \<circ> \<gamma>)) = a" | |
| 2167 | by (simp add: \<open>pathstart qq = a\<close>) | |
| 2168 |             show "p ((qq +++ (p' \<circ> f \<circ> \<gamma>)) \<xi>) = f ((pp +++ \<gamma>) \<xi>)" if \<xi>: "\<xi> \<in> {0..1}" for \<xi>
 | |
| 2169 | proof (simp add: joinpaths_def, safe) | |
| 2170 | show "p (qq (2*\<xi>)) = f (pp (2*\<xi>))" if "\<xi>*2 \<le> 1" | |
| 2171 |                 using \<open>\<xi> \<in> {0..1}\<close> pqqeq that by auto
 | |
| 2172 | show "p (p' (f (\<gamma> (2*\<xi> - 1)))) = f (\<gamma> (2*\<xi> - 1))" if "\<not> \<xi>*2 \<le> 1" | |
| 2173 | apply (rule homeomorphism_apply2 [OF homUW' finW]) | |
| 2174 | using that \<xi> by auto | |
| 2175 | qed | |
| 2176 | qed | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2177 | with \<open>pathfinish \<gamma> = y'\<close> \<open>p' (f y') \<in> X\<close> show "y' \<in> l -` X" | 
| 64792 | 2178 | unfolding pathfinish_join by (simp add: pathfinish_def) | 
| 2179 | qed | |
| 2180 | qed | |
| 2181 | qed | |
| 2182 | then show "continuous_on U l" | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 2183 | by (metis IntD1 IntD2 vimage_eq openin_subopen continuous_on_open_gen [OF LC]) | 
| 64792 | 2184 | qed | 
| 2185 | qed | |
| 2186 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2187 | corollary covering_space_lift_stronger: | 
| 64792 | 2188 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 2189 | and f :: "'c::real_normed_vector \<Rightarrow> 'b" | |
| 2190 | assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U" | |
| 2191 | and U: "path_connected U" "locally path_connected U" | |
| 2192 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S" | |
| 2193 | and feq: "f z = p a" | |
| 2194 | and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk> | |
| 2195 | \<Longrightarrow> \<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)" | |
| 2196 | obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2197 | proof (rule covering_space_lift_general [OF cov U contf fim feq]) | 
| 64792 | 2198 | fix r | 
| 2199 | assume "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z" | |
| 2200 | then obtain b where b: "homotopic_paths S (f \<circ> r) (linepath b b)" | |
| 2201 | using hom by blast | |
| 2202 | then have "f (pathstart r) = b" | |
| 2203 | by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath) | |
| 2204 | then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))" | |
| 2205 | by (simp add: b \<open>pathstart r = z\<close>) | |
| 2206 | then have "homotopic_paths S (f \<circ> r) (p \<circ> linepath a a)" | |
| 2207 | by (simp add: o_def feq linepath_def) | |
| 2208 | then show "\<exists>q. path q \<and> | |
| 2209 | path_image q \<subseteq> C \<and> | |
| 2210 | pathstart q = a \<and> pathfinish q = a \<and> homotopic_paths S (f \<circ> r) (p \<circ> q)" | |
| 2211 | by (force simp: \<open>a \<in> C\<close>) | |
| 2212 | qed auto | |
| 2213 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2214 | corollary covering_space_lift_strong: | 
| 64792 | 2215 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 2216 | and f :: "'c::real_normed_vector \<Rightarrow> 'b" | |
| 2217 | assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U" | |
| 2218 | and scU: "simply_connected U" and lpcU: "locally path_connected U" | |
| 2219 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S" | |
| 2220 | and feq: "f z = p a" | |
| 2221 | obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2222 | proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq]) | 
| 64792 | 2223 | show "path_connected U" | 
| 2224 | using scU simply_connected_eq_contractible_loop_some by blast | |
| 2225 | fix r | |
| 2226 | assume r: "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z" | |
| 2227 | have "linepath (f z) (f z) = f \<circ> linepath z z" | |
| 2228 | by (simp add: o_def linepath_def) | |
| 2229 | then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))" | |
| 2230 | by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path) | |
| 2231 | then show "\<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)" | |
| 2232 | by blast | |
| 2233 | qed blast | |
| 2234 | ||
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2235 | corollary covering_space_lift: | 
| 64792 | 2236 | fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 2237 | and f :: "'c::real_normed_vector \<Rightarrow> 'b" | |
| 2238 | assumes cov: "covering_space C p S" | |
| 2239 | and U: "simply_connected U" "locally path_connected U" | |
| 2240 | and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S" | |
| 2241 | obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y" | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 2242 | proof (cases "U = {}")
 | 
| 64792 | 2243 | case True | 
| 2244 | with that show ?thesis by auto | |
| 2245 | next | |
| 2246 | case False | |
| 2247 | then obtain z where "z \<in> U" by blast | |
| 2248 | then obtain a where "a \<in> C" "f z = p a" | |
| 2249 | by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff) | |
| 2250 | then show ?thesis | |
| 2251 | by (metis that covering_space_lift_strong [OF cov _ \<open>z \<in> U\<close> U contf fim]) | |
| 2252 | qed | |
| 2253 | ||
| 71184 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2254 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Homeomorphisms of arc images\<close> | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2255 | |
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2256 | lemma homeomorphism_arc: | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2257 | fixes g :: "real \<Rightarrow> 'a::t2_space" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2258 | assumes "arc g" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2259 |   obtains h where "homeomorphism {0..1} (path_image g) g h"
 | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2260 | using assms by (force simp: arc_def homeomorphism_compact path_def path_image_def) | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2261 | |
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2262 | lemma homeomorphic_arc_image_interval: | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2263 | fixes g :: "real \<Rightarrow> 'a::t2_space" and a::real | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2264 | assumes "arc g" "a < b" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2265 |   shows "(path_image g) homeomorphic {a..b}"
 | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2266 | proof - | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2267 |   have "(path_image g) homeomorphic {0..1::real}"
 | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2268 | by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2269 |   also have "\<dots> homeomorphic {a..b}"
 | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2270 | using assms by (force intro: homeomorphic_closed_intervals_real) | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2271 | finally show ?thesis . | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2272 | qed | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2273 | |
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2274 | lemma homeomorphic_arc_images: | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2275 | fixes g :: "real \<Rightarrow> 'a::t2_space" and h :: "real \<Rightarrow> 'b::t2_space" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2276 | assumes "arc g" "arc h" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2277 | shows "(path_image g) homeomorphic (path_image h)" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2278 | proof - | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2279 |   have "(path_image g) homeomorphic {0..1::real}"
 | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2280 | by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2281 | also have "\<dots> homeomorphic (path_image h)" | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2282 | by (meson assms homeomorphic_def homeomorphism_arc) | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2283 | finally show ?thesis . | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2284 | qed | 
| 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 2285 | |
| 63130 | 2286 | end |