| author | wenzelm | 
| Fri, 15 Oct 2021 21:25:47 +0200 | |
| changeset 74528 | ce2c7037e509 | 
| parent 71633 | 07bec530f02e | 
| child 78320 | eb9a9690b8f5 | 
| permissions | -rw-r--r-- | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Author: LCP, ported from HOL Light | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | *) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | section\<open>Euclidean space and n-spheres, as subtopologies of n-dimensional space\<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | theory Abstract_Euclidean_Space | 
| 71170 | 7 | imports Homotopy Locally | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | begin | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | subsection \<open>Euclidean spaces as abstract topologies\<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | definition Euclidean_space :: "nat \<Rightarrow> (nat \<Rightarrow> real) topology" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 |   where "Euclidean_space n \<equiv> subtopology (powertop_real UNIV) {x. \<forall>i\<ge>n. x i = 0}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | lemma topspace_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 |    "topspace(Euclidean_space n) = {x. \<forall>i\<ge>n. x i = 0}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | by (simp add: Euclidean_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | lemma nonempty_Euclidean_space: "topspace(Euclidean_space n) \<noteq> {}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | by (force simp: topspace_Euclidean_space) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | lemma subset_Euclidean_space [simp]: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | "topspace(Euclidean_space m) \<subseteq> topspace(Euclidean_space n) \<longleftrightarrow> m \<le> n" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | apply (simp add: topspace_Euclidean_space subset_iff, safe) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | apply (drule_tac x="(\<lambda>i. if i < m then 1 else 0)" in spec) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | apply auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | using not_less by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | lemma topspace_Euclidean_space_alt: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 |   "topspace(Euclidean_space n) = (\<Inter>i \<in> {n..}. {x. x \<in> topspace(powertop_real UNIV) \<and> x i \<in> {0}})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | by (auto simp: topspace_Euclidean_space) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | lemma closedin_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | "closedin (powertop_real UNIV) (topspace(Euclidean_space n))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 |   have "closedin (powertop_real UNIV) {x. x i = 0}" if "n \<le> i" for i
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 |     have "closedin (powertop_real UNIV) {x \<in> topspace (powertop_real UNIV). x i \<in> {0}}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | proof (rule closedin_continuous_map_preimage) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | show "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | by (metis UNIV_I continuous_map_product_coordinates) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 |       show "closedin euclideanreal {0}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | by simp | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | unfolding topspace_Euclidean_space_alt | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | by force | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | lemma closedin_Euclidean_imp_closed: "closedin (Euclidean_space m) S \<Longrightarrow> closed S" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | by (metis Euclidean_space_def closed_closedin closedin_Euclidean_space closedin_closed_subtopology euclidean_product_topology topspace_Euclidean_space) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | lemma closedin_Euclidean_space_iff: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | "closedin (Euclidean_space m) S \<longleftrightarrow> closed S \<and> S \<subseteq> topspace (Euclidean_space m)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | using closedin_closed_subtopology topspace_Euclidean_space | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | by (fastforce simp: topspace_Euclidean_space_alt closedin_Euclidean_imp_closed) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | apply (simp add: closedin_subtopology Euclidean_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | by (metis (no_types) closed_closedin euclidean_product_topology inf.orderE) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | lemma continuous_map_componentwise_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | "continuous_map X (Euclidean_space n) (\<lambda>x i. if i < n then f x i else 0) \<longleftrightarrow> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | (\<forall>i < n. continuous_map X euclideanreal (\<lambda>x. f x i))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | have *: "continuous_map X euclideanreal (\<lambda>x. if k < n then f x k else 0)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | if "\<And>i. i<n \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)" for k | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | by (intro continuous_intros that) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | unfolding Euclidean_space_def continuous_map_in_subtopology | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | by (fastforce simp: continuous_map_componentwise_UNIV * elim: continuous_map_eq) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | lemma continuous_map_Euclidean_space_add [continuous_intros]: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | "\<lbrakk>continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\<rbrakk> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | \<Longrightarrow> continuous_map X (Euclidean_space n) (\<lambda>x i. f x i + g x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | unfolding Euclidean_space_def continuous_map_in_subtopology | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 84 | by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_add) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | lemma continuous_map_Euclidean_space_diff [continuous_intros]: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | "\<lbrakk>continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\<rbrakk> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | \<Longrightarrow> continuous_map X (Euclidean_space n) (\<lambda>x i. f x i - g x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | unfolding Euclidean_space_def continuous_map_in_subtopology | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 90 | by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_diff) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | |
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 92 | lemma continuous_map_Euclidean_space_iff: | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 93 | "continuous_map (Euclidean_space m) euclidean g | 
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 94 | = continuous_on (topspace (Euclidean_space m)) g" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 95 | proof | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 96 | assume "continuous_map (Euclidean_space m) euclidean g" | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 97 |   then have "continuous_map (top_of_set {f. \<forall>n\<ge>m. f n = 0}) euclidean g"
 | 
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 98 | by (simp add: Euclidean_space_def euclidean_product_topology) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 99 | then show "continuous_on (topspace (Euclidean_space m)) g" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 100 | by (metis continuous_map_subtopology_eu subtopology_topspace topspace_Euclidean_space) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 101 | next | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 102 | assume "continuous_on (topspace (Euclidean_space m)) g" | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 103 |   then have "continuous_map (top_of_set {f. \<forall>n\<ge>m. f n = 0}) euclidean g"
 | 
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 104 | by (metis (lifting) continuous_map_into_fulltopology continuous_map_subtopology_eu order_refl topspace_Euclidean_space) | 
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 105 | then show "continuous_map (Euclidean_space m) euclidean g" | 
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 106 | by (simp add: Euclidean_space_def euclidean_product_topology) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 107 | qed | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 108 | |
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 109 | lemma cm_Euclidean_space_iff_continuous_on: | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 110 | "continuous_map (subtopology (Euclidean_space m) S) (Euclidean_space n) f | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 111 | \<longleftrightarrow> continuous_on (topspace (subtopology (Euclidean_space m) S)) f \<and> | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 112 | f ` (topspace (subtopology (Euclidean_space m) S)) \<subseteq> topspace (Euclidean_space n)" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 113 | (is "?P \<longleftrightarrow> ?Q \<and> ?R") | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 114 | proof - | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 115 | have ?Q if ?P | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 116 | proof - | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 117 |     have "\<And>n. Euclidean_space n = top_of_set {f. \<forall>m\<ge>n. f m = 0}"
 | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 118 | by (simp add: Euclidean_space_def euclidean_product_topology) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 119 | with that show ?thesis | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 120 | by (simp add: subtopology_subtopology) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 121 | qed | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 122 | moreover | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 123 | have ?R if ?P | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 124 | using that by (simp add: image_subset_iff continuous_map_def) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 125 | moreover | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 126 | have ?P if ?Q ?R | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 127 | proof - | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 128 |     have "continuous_map (top_of_set (topspace (subtopology (subtopology (powertop_real UNIV) {f. \<forall>n\<ge>m. f n = 0}) S))) (top_of_set (topspace (subtopology (powertop_real UNIV) {f. \<forall>na\<ge>n. f na = 0}))) f"
 | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 129 | using Euclidean_space_def that by auto | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 130 | then show ?thesis | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 131 | by (simp add: Euclidean_space_def euclidean_product_topology subtopology_subtopology) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 132 | qed | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 133 | ultimately show ?thesis | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 134 | by auto | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 135 | qed | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 136 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | lemma homeomorphic_Euclidean_space_product_topology: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 |   "Euclidean_space n homeomorphic_space product_topology (\<lambda>i. euclideanreal) {..<n}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 |   have cm: "continuous_map (product_topology (\<lambda>i. euclideanreal) {..<n})
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | euclideanreal (\<lambda>x. if k < n then x k else 0)" for k | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | by (auto intro: continuous_map_if continuous_map_product_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | unfolding homeomorphic_space_def homeomorphic_maps_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 |     apply (rule_tac x="\<lambda>f. restrict f {..<n}" in exI)
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | apply (rule_tac x="\<lambda>f i. if i < n then f i else 0" in exI) | 
| 71172 | 147 | apply (simp add: Euclidean_space_def continuous_map_in_subtopology) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | apply (intro conjI continuous_map_from_subtopology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | apply (force simp: continuous_map_componentwise cm intro: continuous_map_product_projection)+ | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | lemma contractible_Euclidean_space [simp]: "contractible_space (Euclidean_space n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | using homeomorphic_Euclidean_space_product_topology contractible_space_euclideanreal | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | contractible_space_product_topology homeomorphic_space_contractibility by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | lemma path_connected_Euclidean_space: "path_connected_space (Euclidean_space n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | by (simp add: contractible_imp_path_connected_space) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | lemma connected_Euclidean_space: "connected_space (Euclidean_space n)" | 
| 71633 | 161 | by (simp add: contractible_imp_connected_space) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | lemma locally_path_connected_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | "locally_path_connected_space (Euclidean_space n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | apply (simp add: homeomorphic_locally_path_connected_space [OF homeomorphic_Euclidean_space_product_topology [of n]] | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | locally_path_connected_space_product_topology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | using locally_path_connected_space_euclideanreal by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | lemma compact_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | "compact_space (Euclidean_space n) \<longleftrightarrow> n = 0" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | by (auto simp: homeomorphic_compact_space [OF homeomorphic_Euclidean_space_product_topology] compact_space_product_topology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | |
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 173 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | subsection\<open>n-dimensional spheres\<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | definition nsphere where | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 |  "nsphere n \<equiv> subtopology (Euclidean_space (Suc n)) { x. (\<Sum>i\<le>n. x i ^ 2) = 1 }"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | lemma nsphere: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | "nsphere n = subtopology (powertop_real UNIV) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 |                             {x. (\<Sum>i\<le>n. x i ^ 2) = 1 \<and> (\<forall>i>n. x i = 0)}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | by (simp add: nsphere_def Euclidean_space_def subtopology_subtopology Suc_le_eq Collect_conj_eq Int_commute) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | lemma continuous_map_nsphere_projection: "continuous_map (nsphere n) euclideanreal (\<lambda>x. x k)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | unfolding nsphere | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | by (blast intro: continuous_map_from_subtopology [OF continuous_map_product_projection]) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | lemma in_topspace_nsphere: "(\<lambda>n. if n = 0 then 1 else 0) \<in> topspace (nsphere n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | by (simp add: nsphere_def topspace_Euclidean_space power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | lemma nonempty_nsphere [simp]: "~ (topspace(nsphere n) = {})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | using in_topspace_nsphere by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | lemma subtopology_nsphere_equator: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 |   "subtopology (nsphere (Suc n)) {x. x(Suc n) = 0} = nsphere n"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 |   have "({x. (\<Sum>i\<le>n. (x i)\<^sup>2) + (x (Suc n))\<^sup>2 = 1 \<and> (\<forall>i>Suc n. x i = 0)} \<inter> {x. x (Suc n) = 0})
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 |       = {x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = (0::real))}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | using Suc_lessI [of n] by (fastforce simp: set_eq_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | by (simp add: nsphere subtopology_subtopology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | |
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 204 | lemma topspace_nsphere_minus1: | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 205 | assumes x: "x \<in> topspace (nsphere n)" and "x n = 0" | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 206 | shows "x \<in> topspace (nsphere (n - Suc 0))" | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 207 | proof (cases "n = 0") | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 208 | case True | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 209 | then show ?thesis | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 210 | using x by auto | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 211 | next | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 212 | case False | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 213 |   have subt_eq: "nsphere (n - Suc 0) = subtopology (nsphere n) {x. x n = 0}"
 | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 214 | by (metis False Suc_pred le_zero_eq not_le subtopology_nsphere_equator) | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 215 | with x show ?thesis | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 216 | by (simp add: assms) | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 217 | qed | 
| 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
70086diff
changeset | 218 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | lemma continuous_map_nsphere_reflection: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | "continuous_map (nsphere n) (nsphere n) (\<lambda>x i. if i = k then -x i else x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | proof - | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 222 | have cm: "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. if j = k then - x j else x j)" for j | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | proof (cases "j=k") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | case True | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | then show ?thesis | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 226 | by simp (metis UNIV_I continuous_map_product_projection) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | case False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | by (auto intro: continuous_map_product_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | have eq: "(if i = k then x k * x k else x i * x i) = x i * x i" for i and x :: "nat \<Rightarrow> real" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | by simp | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | apply (simp add: nsphere continuous_map_in_subtopology continuous_map_componentwise_UNIV | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 236 | continuous_map_from_subtopology cm) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | apply (intro conjI allI impI continuous_intros continuous_map_from_subtopology continuous_map_product_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | apply (auto simp: power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] eq cong: if_cong) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | proposition contractible_space_upper_hemisphere: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | assumes "k \<le> n" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 |   shows "contractible_space(subtopology (nsphere n) {x. x k \<ge> 0})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | define p:: "nat \<Rightarrow> real" where "p \<equiv> \<lambda>i. if i = k then 1 else 0" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | have "p \<in> topspace(nsphere n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | using assms | 
| 71172 | 250 | by (simp add: nsphere p_def power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | let ?g = "\<lambda>x i. x i / sqrt(\<Sum>j\<le>n. x j ^ 2)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | let ?h = "\<lambda>(t,q) i. (1 - t) * q i + t * p i" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 |   let ?Y = "subtopology (Euclidean_space (Suc n)) {x. 0 \<le> x k \<and> (\<exists>i\<le>n. x i \<noteq> 0)}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 |   have "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \<le> x k}))
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 |                        (subtopology (nsphere n) {x. 0 \<le> x k}) (?g \<circ> ?h)"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | proof (rule continuous_map_compose) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | have *: "\<lbrakk>0 \<le> b k; (\<Sum>i\<le>n. (b i)\<^sup>2) = 1; \<forall>i>n. b i = 0; 0 \<le> a; a \<le> 1\<rbrakk> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | \<Longrightarrow> \<exists>i. (i = k \<longrightarrow> (1 - a) * b k + a \<noteq> 0) \<and> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | (i \<noteq> k \<longrightarrow> i \<le> n \<and> a \<noteq> 1 \<and> b i \<noteq> 0)" for a::real and b | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | apply (cases "a \<noteq> 1 \<and> b k = 0"; simp) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | apply (metis (no_types, lifting) atMost_iff sum.neutral zero_power2) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by (metis add.commute add_le_same_cancel2 diff_ge_0_iff_ge diff_zero less_eq_real_def mult_eq_0_iff mult_nonneg_nonneg not_le numeral_One zero_neq_numeral) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 |     show "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \<le> x k})) ?Y ?h"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | using assms | 
| 71172 | 265 | apply (auto simp: * nsphere continuous_map_componentwise_UNIV | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | prod_topology_subtopology subtopology_subtopology case_prod_unfold | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | continuous_map_in_subtopology Euclidean_space_def p_def if_distrib [where f = "\<lambda>x. _ * x"] cong: if_cong) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | apply (intro continuous_map_prod_snd continuous_intros continuous_map_from_subtopology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | apply auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | have 1: "\<And>x i. \<lbrakk> i \<le> n; x i \<noteq> 0\<rbrakk> \<Longrightarrow> (\<Sum>i\<le>n. (x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))\<^sup>2) = 1" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70097diff
changeset | 273 | by (force simp: sum_nonneg sum_nonneg_eq_0_iff field_split_simps simp flip: sum_divide_distrib) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | have cm: "continuous_map ?Y (nsphere n) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | unfolding Euclidean_space_def nsphere subtopology_subtopology continuous_map_in_subtopology | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | proof (intro continuous_intros conjI) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | show "continuous_map | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 |                (subtopology (powertop_real UNIV) ({x. \<forall>i\<ge>Suc n. x i = 0} \<inter> {x. 0 \<le> x k \<and> (\<exists>i\<le>n. x i \<noteq> 0)}))
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | (powertop_real UNIV) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | unfolding continuous_map_componentwise | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | by (intro continuous_intros conjI ballI) (auto simp: sum_nonneg_eq_0_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | qed (auto simp: 1) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 |     show "continuous_map ?Y (subtopology (nsphere n) {x. 0 \<le> x k}) (\<lambda>x i. x i / sqrt (\<Sum>j\<le>n. (x j)\<^sup>2))"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | by (force simp: cm sum_nonneg continuous_map_in_subtopology if_distrib [where f = "\<lambda>x. _ * x"] cong: if_cong) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | moreover have "(?g \<circ> ?h) (0, x) = x" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 |     if "x \<in> topspace (subtopology (nsphere n) {x. 0 \<le> x k})" for x
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | using that | 
| 71172 | 289 | by (simp add: assms nsphere) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | moreover | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | have "(?g \<circ> ?h) (1, x) = p" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 |     if "x \<in> topspace (subtopology (nsphere n) {x. 0 \<le> x k})" for x
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | by (force simp: assms p_def power2_eq_square if_distrib [where f = "\<lambda>x. x * _"] cong: if_cong) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | ultimately | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | apply (simp add: contractible_space_def homotopic_with) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | apply (rule_tac x=p in exI) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | apply (rule_tac x="?g \<circ> ?h" in exI, force) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | corollary contractible_space_lower_hemisphere: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | assumes "k \<le> n" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 |   shows "contractible_space(subtopology (nsphere n) {x. x k \<le> 0})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 |   have "contractible_space (subtopology (nsphere n) {x. 0 \<le> x k}) = ?thesis"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | proof (rule homeomorphic_space_contractibility) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 |     show "subtopology (nsphere n) {x. 0 \<le> x k} homeomorphic_space subtopology (nsphere n) {x. x k \<le> 0}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | unfolding homeomorphic_space_def homeomorphic_maps_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | apply (rule_tac x="\<lambda>x i. if i = k then -(x i) else x i" in exI)+ | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | apply (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | continuous_map_nsphere_reflection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | using contractible_space_upper_hemisphere [OF assms] by metis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | proposition nullhomotopic_nonsurjective_sphere_map: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | assumes f: "continuous_map (nsphere p) (nsphere p) f" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | and fim: "f ` (topspace(nsphere p)) \<noteq> topspace(nsphere p)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | obtains a where "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. a)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | obtain a where a: "a \<in> topspace(nsphere p)" "a \<notin> f ` (topspace(nsphere p))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | using fim continuous_map_image_subset_topspace f by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | then have a1: "(\<Sum>i\<le>p. (a i)\<^sup>2) = 1" and a0: "\<And>i. i > p \<Longrightarrow> a i = 0" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | by (simp_all add: nsphere) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | have f1: "(\<Sum>j\<le>p. (f x j)\<^sup>2) = 1" if "x \<in> topspace (nsphere p)" for x | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | have "f x \<in> topspace (nsphere p)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | using continuous_map_image_subset_topspace f that by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | by (simp add: nsphere) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | show thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | let ?g = "\<lambda>x i. x i / sqrt(\<Sum>j\<le>p. x j ^ 2)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | let ?h = "\<lambda>(t,x) i. (1 - t) * f x i - t * a i" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 |     let ?Y = "subtopology (Euclidean_space(Suc p)) (- {\<lambda>i. 0})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 |     let ?T01 = "top_of_set {0..1::real}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | have 1: "continuous_map (prod_topology ?T01 (nsphere p)) (nsphere p) (?g \<circ> ?h)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | proof (rule continuous_map_compose) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal ((\<lambda>x. f x k) \<circ> snd)" for k | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | unfolding nsphere | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | apply (simp add: continuous_map_of_snd) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | apply (rule continuous_map_compose [of _ "nsphere p" f, unfolded o_def]) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | using f apply (simp add: nsphere) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | by (simp add: continuous_map_nsphere_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | then have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal (\<lambda>r. ?h r k)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | for k | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | unfolding case_prod_unfold o_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | by (intro continuous_map_into_fulltopology [OF continuous_map_fst] continuous_intros) auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 |       moreover have "?h ` ({0..1} \<times> topspace (nsphere p)) \<subseteq> {x. \<forall>i\<ge>Suc p. x i = 0}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | using continuous_map_image_subset_topspace [OF f] | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | by (auto simp: nsphere image_subset_iff a0) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 |       moreover have "(\<lambda>i. 0) \<notin> ?h ` ({0..1} \<times> topspace (nsphere p))"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | proof clarify | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | fix t b | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 |         assume eq: "(\<lambda>i. 0) = (\<lambda>i. (1 - t) * f b i - t * a i)" and "t \<in> {0..1}" and b: "b \<in> topspace (nsphere p)"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | have "(1 - t)\<^sup>2 = (\<Sum>i\<le>p. ((1 - t) * f b i)^2)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | using f1 [OF b] by (simp add: power_mult_distrib flip: sum_distrib_left) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | also have "\<dots> = (\<Sum>i\<le>p. (t * a i)^2)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | using eq by (simp add: fun_eq_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | also have "\<dots> = t\<^sup>2" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | using a1 by (simp add: power_mult_distrib flip: sum_distrib_left) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | finally have "1 - t = t" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | by (simp add: power2_eq_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | then have *: "t = 1/2" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | by simp | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | have fba: "f b \<noteq> a" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | using a(2) b by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | then show False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | using eq unfolding * by (simp add: fun_eq_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | ultimately show "continuous_map (prod_topology ?T01 (nsphere p)) ?Y ?h" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | by (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | have *: "\<lbrakk>\<forall>i\<ge>Suc p. x i = 0; x \<noteq> (\<lambda>i. 0)\<rbrakk> \<Longrightarrow> (\<Sum>j\<le>p. (x j)\<^sup>2) \<noteq> 0" for x :: "nat \<Rightarrow> real" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | by (force simp: fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | show "continuous_map ?Y (nsphere p) ?g" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | apply (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | nsphere continuous_map_componentwise subtopology_subtopology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | apply (intro conjI allI continuous_intros continuous_map_from_subtopology [OF continuous_map_product_projection]) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | apply (simp_all add: *) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | apply (force simp: sum_nonneg fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff power_divide simp flip: sum_divide_distrib) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | have 2: "(?g \<circ> ?h) (0, x) = f x" if "x \<in> topspace (nsphere p)" for x | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | using that f1 by simp | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | have 3: "(?g \<circ> ?h) (1, x) = (\<lambda>i. - a i)" for x | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70097diff
changeset | 393 | using a by (force simp: field_split_simps nsphere) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | then show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. (\<lambda>i. - a i))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | by (force simp: homotopic_with intro: 1 2 3) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | lemma Hausdorff_Euclidean_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | "Hausdorff_space (Euclidean_space n)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | unfolding Euclidean_space_def | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 402 | by (rule Hausdorff_space_subtopology) (metis Hausdorff_space_euclidean Hausdorff_space_product_topology) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | end | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 |