| author | wenzelm | 
| Sun, 27 Oct 2024 12:32:40 +0100 | |
| changeset 81275 | 5ed639c16ce7 | 
| parent 80914 | d97fdabd9e2b | 
| child 82323 | b022c013b04b | 
| permissions | -rw-r--r-- | 
| 69620 | 1 | (* Title: HOL/Analysis/Path_Connected.thy | 
| 2 | Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light | |
| 3 | *) | |
| 4 | ||
| 5 | section \<open>Homotopy of Maps\<close> | |
| 6 | ||
| 7 | theory Homotopy | |
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 8 | imports Path_Connected Product_Topology Uncountable_Sets | 
| 69620 | 9 | begin | 
| 10 | ||
| 70136 | 11 | definition\<^marker>\<open>tag important\<close> homotopic_with | 
| 69620 | 12 | where | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 13 | "homotopic_with P X Y f g \<equiv> | 
| 71745 | 14 |    (\<exists>h. continuous_map (prod_topology (top_of_set {0..1::real}) X) Y h \<and>
 | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 15 | (\<forall>x. h(0, x) = f x) \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 16 | (\<forall>x. h(1, x) = g x) \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 17 |        (\<forall>t \<in> {0..1}. P(\<lambda>x. h(t,x))))"
 | 
| 69620 | 18 | |
| 19 | text\<open>\<open>p\<close>, \<open>q\<close> are functions \<open>X \<rightarrow> Y\<close>, and the property \<open>P\<close> restricts all intermediate maps. | |
| 20 | We often just want to require that \<open>P\<close> fixes some subset, but to include the case of a loop homotopy, | |
| 21 | it is convenient to have a general property \<open>P\<close>.\<close> | |
| 22 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 23 | abbreviation homotopic_with_canon :: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 24 |   "[('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool, 'a set, 'b set, 'a \<Rightarrow> 'b, 'a \<Rightarrow> 'b] \<Rightarrow> bool"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 25 | where | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 26 | "homotopic_with_canon P S T p q \<equiv> homotopic_with P (top_of_set S) (top_of_set T) p q" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 27 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 28 | lemma split_01: "{0..1::real} = {0..1/2} \<union> {1/2..1}"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 29 | by force | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 30 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 31 | lemma split_01_prod: "{0..1::real} \<times> X = ({0..1/2} \<times> X) \<union> ({1/2..1} \<times> X)"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 32 | by force | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 33 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 34 | lemma image_Pair_const: "(\<lambda>x. (x, c)) ` A = A \<times> {c}"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 35 | by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 36 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 37 | lemma fst_o_paired [simp]: "fst \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). f x y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 38 | by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 39 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 40 | lemma snd_o_paired [simp]: "snd \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). g x y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 41 | by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 42 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 43 | lemma continuous_on_o_Pair: "\<lbrakk>continuous_on (T \<times> X) h; t \<in> T\<rbrakk> \<Longrightarrow> continuous_on X (h \<circ> Pair t)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 44 | by (fast intro: continuous_intros elim!: continuous_on_subset) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 45 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 46 | lemma continuous_map_o_Pair: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 47 | assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \<in> topspace X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 48 | shows "continuous_map Y Z (h \<circ> Pair t)" | 
| 71745 | 49 | by (intro continuous_map_compose [OF _ h] continuous_intros; simp add: t) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 50 | |
| 70136 | 51 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial properties\<close> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 52 | |
| 69620 | 53 | text \<open>We often want to just localize the ending function equality or whatever.\<close> | 
| 70136 | 54 | text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close> | 
| 69620 | 55 | proposition homotopic_with: | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 56 | assumes "\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k)" | 
| 69620 | 57 | shows "homotopic_with P X Y p q \<longleftrightarrow> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 58 |            (\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and>
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 59 | (\<forall>x \<in> topspace X. h(0,x) = p x) \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 60 | (\<forall>x \<in> topspace X. h(1,x) = q x) \<and> | 
| 69620 | 61 |               (\<forall>t \<in> {0..1}. P(\<lambda>x. h(t, x))))"
 | 
| 62 | unfolding homotopic_with_def | |
| 63 | apply (rule iffI, blast, clarify) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 64 | apply (rule_tac x="\<lambda>(u,v). if v \<in> topspace X then h(u,v) else if u = 0 then p v else q v" in exI) | 
| 77684 | 65 | apply simp | 
| 66 | by (smt (verit, best) SigmaE assms case_prod_conv continuous_map_eq topspace_prod_topology) | |
| 69620 | 67 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 68 | lemma homotopic_with_mono: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 69 | assumes hom: "homotopic_with P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 70 | and Q: "\<And>h. \<lbrakk>continuous_map X Y h; P h\<rbrakk> \<Longrightarrow> Q h" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 71 | shows "homotopic_with Q X Y f g" | 
| 71745 | 72 | using hom unfolding homotopic_with_def | 
| 73 | by (force simp: o_def dest: continuous_map_o_Pair intro: Q) | |
| 69620 | 74 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 75 | lemma homotopic_with_imp_continuous_maps: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 76 | assumes "homotopic_with P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 77 | shows "continuous_map X Y f \<and> continuous_map X Y g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 78 | proof - | 
| 71745 | 79 | obtain h :: "real \<times> 'a \<Rightarrow> 'b" | 
| 80 |     where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) Y h"
 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 81 | and h: "\<forall>x. h (0, x) = f x" "\<forall>x. h (1, x) = g x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 82 | using assms by (auto simp: homotopic_with_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 83 |   have *: "t \<in> {0..1} \<Longrightarrow> continuous_map X Y (h \<circ> (\<lambda>x. (t,x)))" for t
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 84 | by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 85 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 86 | using h *[of 0] *[of 1] by (simp add: continuous_map_eq) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 87 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 88 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 89 | lemma homotopic_with_imp_continuous: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 90 | assumes "homotopic_with_canon P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 91 | shows "continuous_on X f \<and> continuous_on X g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 92 | by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 93 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 94 | lemma homotopic_with_imp_property: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 95 | assumes "homotopic_with P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 96 | shows "P f \<and> P g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 97 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 98 |   obtain h where h: "\<And>x. h(0, x) = f x" "\<And>x. h(1, x) = g x" and P: "\<And>t. t \<in> {0..1::real} \<Longrightarrow> P(\<lambda>x. h(t,x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 99 | using assms by (force simp: homotopic_with_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 100 | show "P f" "P g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 101 | using P [of 0] P [of 1] by (force simp: h)+ | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 102 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 103 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 104 | lemma homotopic_with_equal: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 105 | assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 106 | shows "homotopic_with P X Y f g" | 
| 69620 | 107 | unfolding homotopic_with_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 108 | proof (intro exI conjI allI ballI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 109 | let ?h = "\<lambda>(t::real,x). if t = 1 then g x else f x" | 
| 71745 | 110 |   show "continuous_map (prod_topology (top_of_set {0..1}) X) Y ?h"
 | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 111 | proof (rule continuous_map_eq) | 
| 71745 | 112 |     show "continuous_map (prod_topology (top_of_set {0..1}) X) Y (f \<circ> snd)"
 | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 113 | by (simp add: contf continuous_map_of_snd) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 114 | qed (auto simp: fg) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 115 |   show "P (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 116 | by (cases "t = 1") (simp_all add: assms) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 117 | qed auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 118 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 119 | lemma homotopic_with_imp_subset1: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 120 | "homotopic_with_canon P X Y f g \<Longrightarrow> f ` X \<subseteq> Y" | 
| 77684 | 121 | by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 122 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 123 | lemma homotopic_with_imp_subset2: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 124 | "homotopic_with_canon P X Y f g \<Longrightarrow> g ` X \<subseteq> Y" | 
| 77684 | 125 | by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 126 | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 127 | lemma homotopic_with_imp_funspace1: | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 128 | "homotopic_with_canon P X Y f g \<Longrightarrow> f \<in> X \<rightarrow> Y" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 129 | using homotopic_with_imp_subset1 by blast | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 130 | |
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 131 | lemma homotopic_with_imp_funspace2: | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 132 | "homotopic_with_canon P X Y f g \<Longrightarrow> g \<in> X \<rightarrow> Y" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 133 | using homotopic_with_imp_subset2 by blast | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 134 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 135 | lemma homotopic_with_subset_left: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 136 | "\<lbrakk>homotopic_with_canon P X Y f g; Z \<subseteq> X\<rbrakk> \<Longrightarrow> homotopic_with_canon P Z Y f g" | 
| 71745 | 137 | unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 138 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 139 | lemma homotopic_with_subset_right: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 140 | "\<lbrakk>homotopic_with_canon P X Y f g; Y \<subseteq> Z\<rbrakk> \<Longrightarrow> homotopic_with_canon P X Z f g" | 
| 71745 | 141 | unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) | 
| 69620 | 142 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 143 | subsection\<open>Homotopy with P is an equivalence relation\<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 144 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 145 | text \<open>(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 146 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 147 | lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \<longleftrightarrow> continuous_map X Y f \<and> P f" | 
| 77684 | 148 | by (metis homotopic_with_equal homotopic_with_imp_continuous_maps homotopic_with_imp_property) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 149 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 150 | lemma homotopic_with_symD: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 151 | assumes "homotopic_with P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 152 | shows "homotopic_with P X Y g f" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 153 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 154 |   let ?I01 = "subtopology euclideanreal {0..1}"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 155 | let ?j = "\<lambda>y. (1 - fst y, snd y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 156 | have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" | 
| 71745 | 157 | by (intro continuous_intros; simp add: continuous_map_subtopology_fst prod_topology_subtopology) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 158 | have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 159 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 160 |     have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \<times> topspace X)) ?j"
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 161 | by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff flip: image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 162 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 163 | by (simp add: prod_topology_subtopology(1)) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 164 | qed | 
| 69620 | 165 | show ?thesis | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 166 | using assms | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 167 | apply (clarsimp simp: homotopic_with_def) | 
| 71745 | 168 | subgoal for h | 
| 169 | by (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *]) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 170 | done | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 171 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 172 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 173 | lemma homotopic_with_sym: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 174 | "homotopic_with P X Y f g \<longleftrightarrow> homotopic_with P X Y g f" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 175 | by (metis homotopic_with_symD) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 176 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 177 | proposition homotopic_with_trans: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 178 | assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 179 | shows "homotopic_with P X Y f h" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 180 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 181 |   let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 182 | obtain k1 k2 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 183 | where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 184 | and k12: "\<forall>x. k1 (1, x) = g x" "\<forall>x. k2 (0, x) = g x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 185 | "\<forall>x. k1 (0, x) = f x" "\<forall>x. k2 (1, x) = h x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 186 |       and P:   "\<forall>t\<in>{0..1}. P (\<lambda>x. k1 (t, x))" "\<forall>t\<in>{0..1}. P (\<lambda>x. k2 (t, x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 187 | using assms by (auto simp: homotopic_with_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 188 | define k where "k \<equiv> \<lambda>y. if fst y \<le> 1/2 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 189 | then (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 190 | else (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 191 | have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 192 | by (simp add: k12 that) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 193 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 194 | unfolding homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 195 | proof (intro exI conjI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 196 | show "continuous_map ?X01 Y k" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 197 | unfolding k_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 198 | proof (rule continuous_map_cases_le) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 199 | show fst: "continuous_map ?X01 euclideanreal fst" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 200 | using continuous_map_fst continuous_map_in_subtopology by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 201 | show "continuous_map ?X01 euclideanreal (\<lambda>x. 1/2)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 202 | by simp | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 203 |       show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. fst y \<le> 1/2}) Y
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 204 | (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x)))" | 
| 72372 | 205 | apply (intro fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ | 
| 71745 | 206 | by (force simp: prod_topology_subtopology) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 207 |       show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. 1/2 \<le> fst y}) Y
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 208 | (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x)))" | 
| 72372 | 209 | apply (intro fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ | 
| 71745 | 210 | by (force simp: prod_topology_subtopology) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 211 | show "(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y = (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 212 | if "y \<in> topspace ?X01" and "fst y = 1/2" for y | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 213 | using that by (simp add: keq) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 214 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 215 | show "\<forall>x. k (0, x) = f x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 216 | by (simp add: k12 k_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 217 | show "\<forall>x. k (1, x) = h x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 218 | by (simp add: k12 k_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 219 |     show "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))"
 | 
| 71745 | 220 | proof | 
| 221 |       fix t show "t\<in>{0..1} \<Longrightarrow> P (\<lambda>x. k (t, x))"
 | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 222 | by (cases "t \<le> 1/2") (auto simp: k_def P) | 
| 71745 | 223 | qed | 
| 69620 | 224 | qed | 
| 225 | qed | |
| 226 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 227 | lemma homotopic_with_id2: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 228 | "(\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x) \<Longrightarrow> homotopic_with (\<lambda>x. True) X X (g \<circ> f) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 229 | by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 230 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 231 | subsection\<open>Continuity lemmas\<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 232 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 233 | lemma homotopic_with_compose_continuous_map_left: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 234 | "\<lbrakk>homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \<And>j. p j \<Longrightarrow> q(h \<circ> j)\<rbrakk> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 235 | \<Longrightarrow> homotopic_with q X1 X3 (h \<circ> f) (h \<circ> g)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 236 | unfolding homotopic_with_def | 
| 69620 | 237 | apply clarify | 
| 71745 | 238 | subgoal for k | 
| 239 | by (rule_tac x="h \<circ> k" in exI) (rule conjI continuous_map_compose | simp add: o_def)+ | |
| 69620 | 240 | done | 
| 241 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 242 | lemma homotopic_with_compose_continuous_map_right: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 243 | assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 244 | and q: "\<And>j. p j \<Longrightarrow> q(j \<circ> h)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 245 | shows "homotopic_with q X1 X3 (f \<circ> h) (g \<circ> h)" | 
| 69620 | 246 | proof - | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 247 | obtain k | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 248 |     where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 249 |       and k: "\<forall>x. k (0, x) = f x" "\<forall>x. k (1, x) = g x" and p: "\<And>t. t\<in>{0..1} \<Longrightarrow> p (\<lambda>x. k (t, x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 250 | using hom unfolding homotopic_with_def by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 251 |   have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \<circ> snd)"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 252 | by (rule continuous_map_compose [OF continuous_map_snd conth]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 253 | let ?h = "k \<circ> (\<lambda>(t,x). (t,h x))" | 
| 69620 | 254 | show ?thesis | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 255 | unfolding homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 256 | proof (intro exI conjI allI ballI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 257 |     have "continuous_map (prod_topology (top_of_set {0..1}) X1)
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 258 |      (prod_topology (top_of_set {0..1::real}) X2) (\<lambda>(t, x). (t, h x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 259 | by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 260 |     then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 261 | by (intro conjI continuous_map_compose [OF _ contk]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 262 |     show "q (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 263 | using q [OF p [OF that]] by (simp add: o_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 264 | qed (auto simp: k) | 
| 69620 | 265 | qed | 
| 266 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 267 | corollary homotopic_compose: | 
| 71745 | 268 | assumes "homotopic_with (\<lambda>x. True) X Y f f'" "homotopic_with (\<lambda>x. True) Y Z g g'" | 
| 269 | shows "homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g' \<circ> f')" | |
| 77684 | 270 | by (metis assms homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right homotopic_with_imp_continuous_maps homotopic_with_trans) | 
| 69620 | 271 | |
| 272 | proposition homotopic_with_compose_continuous_right: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 273 | "\<lbrakk>homotopic_with_canon (\<lambda>f. p (f \<circ> h)) X Y f g; continuous_on W h; h \<in> W \<rightarrow> X\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 274 | \<Longrightarrow> homotopic_with_canon p W Y (f \<circ> h) (g \<circ> h)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 275 | by (simp add: homotopic_with_compose_continuous_map_right image_subset_iff_funcset) | 
| 69620 | 276 | |
| 277 | proposition homotopic_with_compose_continuous_left: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 278 | "\<lbrakk>homotopic_with_canon (\<lambda>f. p (h \<circ> f)) X Y f g; continuous_on Y h; h \<in> Y \<rightarrow> Z\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 279 | \<Longrightarrow> homotopic_with_canon p X Z (h \<circ> f) (h \<circ> g)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 280 | by (simp add: homotopic_with_compose_continuous_map_left image_subset_iff_funcset) | 
| 69620 | 281 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 282 | lemma homotopic_from_subtopology: | 
| 77684 | 283 | "homotopic_with P X X' f g \<Longrightarrow> homotopic_with P (subtopology X S) X' f g" | 
| 284 | by (metis continuous_map_id_subt homotopic_with_compose_continuous_map_right o_id) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 285 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 286 | lemma homotopic_on_emptyI: | 
| 78336 | 287 | assumes "P f" "P g" | 
| 288 | shows "homotopic_with P trivial_topology X f g" | |
| 289 | by (metis assms continuous_map_on_empty empty_iff homotopic_with_equal topspace_discrete_topology) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 290 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 291 | lemma homotopic_on_empty: | 
| 78336 | 292 | "(homotopic_with P trivial_topology X f g \<longleftrightarrow> P f \<and> P g)" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 293 | using homotopic_on_emptyI homotopic_with_imp_property by metis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 294 | |
| 78336 | 295 | lemma homotopic_with_canon_on_empty: "homotopic_with_canon (\<lambda>x. True) {} t f g"
 | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 296 | by (auto intro: homotopic_with_equal) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 297 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 298 | lemma homotopic_constant_maps: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 299 | "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b) \<longleftrightarrow> | 
| 78336 | 300 | X = trivial_topology \<or> path_component_of X' a b" (is "?lhs = ?rhs") | 
| 301 | proof (cases "X = trivial_topology") | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 302 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 303 | then obtain c where c: "c \<in> topspace X" | 
| 78336 | 304 | by fastforce | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 305 |   have "\<exists>g. continuous_map (top_of_set {0..1::real}) X' g \<and> g 0 = a \<and> g 1 = b"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 306 | if "x \<in> topspace X" and hom: "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b)" for x | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 307 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 308 | obtain h :: "real \<times> 'a \<Rightarrow> 'b" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 309 |       where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 310 | and h: "\<And>x. h (0, x) = a" "\<And>x. h (1, x) = b" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 311 | using hom by (auto simp: homotopic_with_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 312 |     have cont: "continuous_map (top_of_set {0..1}) X' (h \<circ> (\<lambda>t. (t, c)))"
 | 
| 78336 | 313 | by (rule continuous_map_compose [OF _ conth] continuous_intros | simp add: c)+ | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 314 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 315 | by (force simp: h) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 316 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 317 | moreover have "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. g 0) (\<lambda>x. g 1)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 318 |     if "x \<in> topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 319 | for x and g :: "real \<Rightarrow> 'b" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 320 | unfolding homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 321 | by (force intro!: continuous_map_compose continuous_intros c that) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 322 | ultimately show ?thesis | 
| 78336 | 323 | using False | 
| 324 | by (metis c path_component_of_set pathin_def) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 325 | qed (simp add: homotopic_on_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 326 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 327 | proposition homotopic_with_eq: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 328 | assumes h: "homotopic_with P X Y f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 329 | and f': "\<And>x. x \<in> topspace X \<Longrightarrow> f' x = f x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 330 | and g': "\<And>x. x \<in> topspace X \<Longrightarrow> g' x = g x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 331 | and P: "(\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> P h \<longleftrightarrow> P k)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 332 | shows "homotopic_with P X Y f' g'" | 
| 77684 | 333 | by (smt (verit, ccfv_SIG) assms homotopic_with) | 
| 69620 | 334 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 335 | lemma homotopic_with_prod_topology: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 336 | assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 337 | and r: "\<And>i j. \<lbrakk>p i; q j\<rbrakk> \<Longrightarrow> r(\<lambda>(x,y). (i x, j y))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 338 | shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 339 | (\<lambda>z. (f(fst z),g(snd z))) (\<lambda>z. (f'(fst z), g'(snd z)))" | 
| 69620 | 340 | proof - | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 341 | obtain h | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 342 |     where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 343 | and h0: "\<And>x. h (0, x) = f x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 344 | and h1: "\<And>x. h (1, x) = f' x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 345 | and p: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p (\<lambda>x. h (t,x))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 346 | using assms unfolding homotopic_with_def by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 347 | obtain k | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 348 |     where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 349 | and k0: "\<And>x. k (0, x) = g x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 350 | and k1: "\<And>x. k (1, x) = g' x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 351 | and q: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> q (\<lambda>x. k (t,x))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 352 | using assms unfolding homotopic_with_def by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 353 | let ?hk = "\<lambda>(t,x,y). (h(t,x), k(t,y))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 354 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 355 | unfolding homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 356 | proof (intro conjI allI exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 357 |     show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2))
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 358 | (prod_topology Y1 Y2) ?hk" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 359 | unfolding continuous_map_pairwise case_prod_unfold | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 360 | by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 361 | continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 362 | continuous_map_compose [OF _ h, unfolded o_def] | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 363 | continuous_map_compose [OF _ k, unfolded o_def])+ | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 364 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 365 | fix x | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 366 | show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 367 | by (auto simp: case_prod_beta h0 k0 h1 k1) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 368 | qed (auto simp: p q r) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 369 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 370 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 371 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 372 | lemma homotopic_with_product_topology: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 373 | assumes ht: "\<And>i. i \<in> I \<Longrightarrow> homotopic_with (p i) (X i) (Y i) (f i) (g i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 374 | and pq: "\<And>h. (\<And>i. i \<in> I \<Longrightarrow> p i (h i)) \<Longrightarrow> q(\<lambda>x. (\<lambda>i\<in>I. h i (x i)))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 375 | shows "homotopic_with q (product_topology X I) (product_topology Y I) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 376 | (\<lambda>z. (\<lambda>i\<in>I. (f i) (z i))) (\<lambda>z. (\<lambda>i\<in>I. (g i) (z i)))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 377 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 378 | obtain h | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 379 |     where h: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 380 | and h0: "\<And>i x. i \<in> I \<Longrightarrow> h i (0, x) = f i x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 381 | and h1: "\<And>i x. i \<in> I \<Longrightarrow> h i (1, x) = g i x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 382 |       and p: "\<And>i t. \<lbrakk>i \<in> I; t \<in> {0..1}\<rbrakk> \<Longrightarrow> p i (\<lambda>x. h i (t,x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 383 | using ht unfolding homotopic_with_def by metis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 384 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 385 | unfolding homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 386 | proof (intro conjI allI exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 387 | let ?h = "\<lambda>(t,z). \<lambda>i\<in>I. h i (t,z i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 388 |     have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 389 | (Y i) (\<lambda>x. h i (fst x, snd x i))" if "i \<in> I" for i | 
| 71745 | 390 | proof - | 
| 391 |       have \<section>: "continuous_map (prod_topology (top_of_set {0..1}) (product_topology X I)) (X i) (\<lambda>x. snd x i)"
 | |
| 392 | using continuous_map_componentwise continuous_map_snd that by fastforce | |
| 393 | show ?thesis | |
| 394 | unfolding continuous_map_pairwise case_prod_unfold | |
| 395 | by (intro conjI that \<section> continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) | |
| 396 | qed | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 397 |     then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I))
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 398 | (product_topology Y I) ?h" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 399 | by (auto simp: continuous_map_componentwise case_prod_beta) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 400 | show "?h (0, x) = (\<lambda>i\<in>I. f i (x i))" "?h (1, x) = (\<lambda>i\<in>I. g i (x i))" for x | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 401 | by (auto simp: case_prod_beta h0 h1) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 402 |     show "\<forall>t\<in>{0..1}. q (\<lambda>x. ?h (t, x))"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 403 | by (force intro: p pq) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 404 | qed | 
| 69620 | 405 | qed | 
| 406 | ||
| 407 | text\<open>Homotopic triviality implicitly incorporates path-connectedness.\<close> | |
| 408 | lemma homotopic_triviality: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 409 | shows "(\<forall>f g. continuous_on S f \<and> f \<in> S \<rightarrow> T \<and> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 410 | continuous_on S g \<and> g \<in> S \<rightarrow> T | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 411 | \<longrightarrow> homotopic_with_canon (\<lambda>x. True) S T f g) \<longleftrightarrow> | 
| 69620 | 412 |           (S = {} \<or> path_connected T) \<and>
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 413 | (\<forall>f. continuous_on S f \<and> f \<in> S \<rightarrow> T \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)))" | 
| 69620 | 414 | (is "?lhs = ?rhs") | 
| 415 | proof (cases "S = {} \<or> T = {}")
 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 416 | case True then show ?thesis | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 417 | by (auto simp: homotopic_on_emptyI simp flip: image_subset_iff_funcset) | 
| 69620 | 418 | next | 
| 419 | case False show ?thesis | |
| 420 | proof | |
| 421 | assume LHS [rule_format]: ?lhs | |
| 422 | have pab: "path_component T a b" if "a \<in> T" "b \<in> T" for a b | |
| 423 | proof - | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 424 | have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. b)" | 
| 71172 | 425 | by (simp add: LHS image_subset_iff that) | 
| 69620 | 426 | then show ?thesis | 
| 78336 | 427 | using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] | 
| 428 | by (metis path_component_of_canon_iff topspace_discrete_topology topspace_euclidean_subtopology) | |
| 69620 | 429 | qed | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 430 | moreover | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 431 | have "\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" if "continuous_on S f" "f \<in> S \<rightarrow> T" for f | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 432 | using False LHS continuous_on_const that by blast | 
| 69620 | 433 | ultimately show ?rhs | 
| 434 | by (simp add: path_connected_component) | |
| 435 | next | |
| 436 | assume RHS: ?rhs | |
| 437 | with False have T: "path_connected T" | |
| 438 | by blast | |
| 439 | show ?lhs | |
| 440 | proof clarify | |
| 441 | fix f g | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 442 | assume "continuous_on S f" "f \<in> S \<rightarrow> T" "continuous_on S g" "g \<in> S \<rightarrow> T" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 443 | obtain c d where c: "homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" and d: "homotopic_with_canon (\<lambda>x. True) S T g (\<lambda>x. d)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 444 | using RHS \<open>continuous_on S f\<close> \<open>continuous_on S g\<close> \<open>f \<in> S \<rightarrow> T\<close> \<open>g \<in> S \<rightarrow> T\<close> by presburger | 
| 69620 | 445 | with T have "path_component T c d" | 
| 77684 | 446 | by (metis False ex_in_conv homotopic_with_imp_subset2 image_subset_iff path_connected_component) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 447 | then have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. c) (\<lambda>x. d)" | 
| 69620 | 448 | by (simp add: homotopic_constant_maps) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 449 | with c d show "homotopic_with_canon (\<lambda>x. True) S T f g" | 
| 69620 | 450 | by (meson homotopic_with_symD homotopic_with_trans) | 
| 451 | qed | |
| 452 | qed | |
| 453 | qed | |
| 454 | ||
| 455 | ||
| 456 | subsection\<open>Homotopy of paths, maintaining the same endpoints\<close> | |
| 457 | ||
| 458 | ||
| 70136 | 459 | definition\<^marker>\<open>tag important\<close> homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool" | 
| 69620 | 460 | where | 
| 77684 | 461 | "homotopic_paths S p q \<equiv> | 
| 462 |        homotopic_with_canon (\<lambda>r. pathstart r = pathstart p \<and> pathfinish r = pathfinish p) {0..1} S p q"
 | |
| 69620 | 463 | |
| 464 | lemma homotopic_paths: | |
| 77684 | 465 | "homotopic_paths S p q \<longleftrightarrow> | 
| 69620 | 466 |       (\<exists>h. continuous_on ({0..1} \<times> {0..1}) h \<and>
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 467 |           h \<in> ({0..1} \<times> {0..1}) \<rightarrow> S \<and>
 | 
| 69620 | 468 |           (\<forall>x \<in> {0..1}. h(0,x) = p x) \<and>
 | 
| 469 |           (\<forall>x \<in> {0..1}. h(1,x) = q x) \<and>
 | |
| 470 |           (\<forall>t \<in> {0..1::real}. pathstart(h \<circ> Pair t) = pathstart p \<and>
 | |
| 471 | pathfinish(h \<circ> Pair t) = pathfinish p))" | |
| 472 | by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) | |
| 473 | ||
| 474 | proposition homotopic_paths_imp_pathstart: | |
| 77684 | 475 | "homotopic_paths S p q \<Longrightarrow> pathstart p = pathstart q" | 
| 69620 | 476 | by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) | 
| 477 | ||
| 478 | proposition homotopic_paths_imp_pathfinish: | |
| 77684 | 479 | "homotopic_paths S p q \<Longrightarrow> pathfinish p = pathfinish q" | 
| 69620 | 480 | by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) | 
| 481 | ||
| 482 | lemma homotopic_paths_imp_path: | |
| 77684 | 483 | "homotopic_paths S p q \<Longrightarrow> path p \<and> path q" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 484 | using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast | 
| 69620 | 485 | |
| 486 | lemma homotopic_paths_imp_subset: | |
| 77684 | 487 | "homotopic_paths S p q \<Longrightarrow> path_image p \<subseteq> S \<and> path_image q \<subseteq> S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 488 | by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps path_image_def) | 
| 69620 | 489 | |
| 77684 | 490 | proposition homotopic_paths_refl [simp]: "homotopic_paths S p p \<longleftrightarrow> path p \<and> path_image p \<subseteq> S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 491 | by (simp add: homotopic_paths_def path_def path_image_def) | 
| 69620 | 492 | |
| 77684 | 493 | proposition homotopic_paths_sym: "homotopic_paths S p q \<Longrightarrow> homotopic_paths S q p" | 
| 69620 | 494 | by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) | 
| 495 | ||
| 77684 | 496 | proposition homotopic_paths_sym_eq: "homotopic_paths S p q \<longleftrightarrow> homotopic_paths S q p" | 
| 69620 | 497 | by (metis homotopic_paths_sym) | 
| 498 | ||
| 499 | proposition homotopic_paths_trans [trans]: | |
| 77684 | 500 | assumes "homotopic_paths S p q" "homotopic_paths S q r" | 
| 501 | shows "homotopic_paths S p r" | |
| 502 | using assms homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart unfolding homotopic_paths_def | |
| 503 | by (smt (verit, ccfv_SIG) homotopic_with_mono homotopic_with_trans) | |
| 69620 | 504 | |
| 505 | proposition homotopic_paths_eq: | |
| 77684 | 506 |      "\<lbrakk>path p; path_image p \<subseteq> S; \<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t\<rbrakk> \<Longrightarrow> homotopic_paths S p q"
 | 
| 507 | by (smt (verit, best) homotopic_paths homotopic_paths_refl) | |
| 69620 | 508 | |
| 509 | proposition homotopic_paths_reparametrize: | |
| 510 | assumes "path p" | |
| 77684 | 511 | and pips: "path_image p \<subseteq> S" | 
| 69620 | 512 |       and contf: "continuous_on {0..1} f"
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 513 |       and f01 :"f \<in> {0..1} \<rightarrow> {0..1}"
 | 
| 69620 | 514 | and [simp]: "f(0) = 0" "f(1) = 1" | 
| 515 |       and q: "\<And>t. t \<in> {0..1} \<Longrightarrow> q(t) = p(f t)"
 | |
| 77684 | 516 | shows "homotopic_paths S p q" | 
| 69620 | 517 | proof - | 
| 518 |   have contp: "continuous_on {0..1} p"
 | |
| 519 | by (metis \<open>path p\<close> path_def) | |
| 520 |   then have "continuous_on {0..1} (p \<circ> f)"
 | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 521 | by (meson assms(4) contf continuous_on_compose continuous_on_subset image_subset_iff_funcset) | 
| 69620 | 522 | then have "path q" | 
| 523 | by (simp add: path_def) (metis q continuous_on_cong) | |
| 77684 | 524 | have piqs: "path_image q \<subseteq> S" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 525 | by (smt (verit, ccfv_threshold) Pi_iff assms(2) assms(4) assms(7) image_subset_iff path_defs(4)) | 
| 69620 | 526 | have fb0: "\<And>a b. \<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> (1 - a) * f b + a * b" | 
| 527 | using f01 by force | |
| 528 | have fb1: "\<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> (1 - a) * f b + a * b \<le> 1" for a b | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 529 | by (intro convex_bound_le) (use f01 in auto) | 
| 77684 | 530 | have "homotopic_paths S q p" | 
| 69620 | 531 | proof (rule homotopic_paths_trans) | 
| 77684 | 532 | show "homotopic_paths S q (p \<circ> f)" | 
| 69620 | 533 | using q by (force intro: homotopic_paths_eq [OF \<open>path q\<close> piqs]) | 
| 534 | next | |
| 77684 | 535 | show "homotopic_paths S (p \<circ> f) p" | 
| 72372 | 536 | using pips [unfolded path_image_def] | 
| 69620 | 537 | apply (simp add: homotopic_paths_def homotopic_with_def) | 
| 538 | apply (rule_tac x="p \<circ> (\<lambda>y. (1 - (fst y)) *\<^sub>R ((f \<circ> snd) y) + (fst y) *\<^sub>R snd y)" in exI) | |
| 539 | apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ | |
| 72372 | 540 | by (auto simp: fb0 fb1 pathstart_def pathfinish_def) | 
| 69620 | 541 | qed | 
| 542 | then show ?thesis | |
| 543 | by (simp add: homotopic_paths_sym) | |
| 544 | qed | |
| 545 | ||
| 77684 | 546 | lemma homotopic_paths_subset: "\<lbrakk>homotopic_paths S p q; S \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t p q" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 547 | unfolding homotopic_paths by fast | 
| 69620 | 548 | |
| 549 | ||
| 550 | text\<open> A slightly ad-hoc but useful lemma in constructing homotopies.\<close> | |
| 71746 | 551 | lemma continuous_on_homotopic_join_lemma: | 
| 69620 | 552 | fixes q :: "[real,real] \<Rightarrow> 'a::topological_space" | 
| 71745 | 553 |   assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))" (is "continuous_on ?A ?p")
 | 
| 554 |       and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))" (is "continuous_on ?A ?q")
 | |
| 69620 | 555 |       and pf: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish(p t) = pathstart(q t)"
 | 
| 556 |     shows "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. (p(fst y) +++ q(fst y)) (snd y))"
 | |
| 557 | proof - | |
| 71745 | 558 | have \<section>: "(\<lambda>t. p (fst t) (2 * snd t)) = ?p \<circ> (\<lambda>y. (fst y, 2 * snd y))" | 
| 559 | "(\<lambda>t. q (fst t) (2 * snd t - 1)) = ?q \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))" | |
| 560 | by force+ | |
| 69620 | 561 | show ?thesis | 
| 71745 | 562 | unfolding joinpaths_def | 
| 563 | proof (rule continuous_on_cases_le) | |
| 564 |     show "continuous_on {y \<in> ?A. snd y \<le> 1/2} (\<lambda>t. p (fst t) (2 * snd t))" 
 | |
| 565 |          "continuous_on {y \<in> ?A. 1/2 \<le> snd y} (\<lambda>t. q (fst t) (2 * snd t - 1))"
 | |
| 566 | "continuous_on ?A snd" | |
| 567 | unfolding \<section> | |
| 568 | by (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ | |
| 569 | qed (use pf in \<open>auto simp: mult.commute pathstart_def pathfinish_def\<close>) | |
| 69620 | 570 | qed | 
| 571 | ||
| 572 | text\<open> Congruence properties of homotopy w.r.t. path-combining operations.\<close> | |
| 573 | ||
| 574 | lemma homotopic_paths_reversepath_D: | |
| 77684 | 575 | assumes "homotopic_paths S p q" | 
| 576 | shows "homotopic_paths S (reversepath p) (reversepath q)" | |
| 69620 | 577 | using assms | 
| 578 | apply (simp add: homotopic_paths_def homotopic_with_def, clarify) | |
| 579 | apply (rule_tac x="h \<circ> (\<lambda>x. (fst x, 1 - snd x))" in exI) | |
| 580 | apply (rule conjI continuous_intros)+ | |
| 581 | apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) | |
| 582 | done | |
| 583 | ||
| 584 | proposition homotopic_paths_reversepath: | |
| 77684 | 585 | "homotopic_paths S (reversepath p) (reversepath q) \<longleftrightarrow> homotopic_paths S p q" | 
| 69620 | 586 | using homotopic_paths_reversepath_D by force | 
| 587 | ||
| 588 | ||
| 589 | proposition homotopic_paths_join: | |
| 77684 | 590 | "\<lbrakk>homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths S (p +++ q) (p' +++ q')" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 591 | apply (clarsimp simp: homotopic_paths_def homotopic_with_def) | 
| 69620 | 592 | apply (rename_tac k1 k2) | 
| 593 | apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI) | |
| 71746 | 594 | apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def) | 
| 69620 | 595 | done | 
| 596 | ||
| 597 | proposition homotopic_paths_continuous_image: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 598 | "\<lbrakk>homotopic_paths S f g; continuous_on S h; h \<in> S \<rightarrow> t\<rbrakk> \<Longrightarrow> homotopic_paths t (h \<circ> f) (h \<circ> g)" | 
| 69620 | 599 | unfolding homotopic_paths_def | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 600 | by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose image_subset_iff_funcset) | 
| 69620 | 601 | |
| 602 | ||
| 603 | subsection\<open>Group properties for homotopy of paths\<close> | |
| 604 | ||
| 70136 | 605 | text\<^marker>\<open>tag important\<close>\<open>So taking equivalence classes under homotopy would give the fundamental group\<close> | 
| 69620 | 606 | |
| 607 | proposition homotopic_paths_rid: | |
| 77684 | 608 | assumes "path p" "path_image p \<subseteq> S" | 
| 609 | shows "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) p" | |
| 71745 | 610 | proof - | 
| 611 |   have \<section>: "continuous_on {0..1} (\<lambda>t::real. if t \<le> 1/2 then 2 *\<^sub>R t else 1)"
 | |
| 612 | unfolding split_01 | |
| 613 | by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ | |
| 614 | show ?thesis | |
| 72372 | 615 | apply (rule homotopic_paths_sym) | 
| 616 | using assms unfolding pathfinish_def joinpaths_def | |
| 617 | by (intro \<section> continuous_on_cases continuous_intros homotopic_paths_reparametrize [where f = "\<lambda>t. if t \<le> 1/2 then 2 *\<^sub>R t else 1"]; force) | |
| 71745 | 618 | qed | 
| 69620 | 619 | |
| 620 | proposition homotopic_paths_lid: | |
| 77684 | 621 | "\<lbrakk>path p; path_image p \<subseteq> S\<rbrakk> \<Longrightarrow> homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p) p" | 
| 622 | using homotopic_paths_rid [of "reversepath p" S] | |
| 69620 | 623 | by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath | 
| 624 | pathfinish_reversepath reversepath_joinpaths reversepath_linepath) | |
| 625 | ||
| 80052 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 626 | lemma homotopic_paths_rid': | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 627 | assumes "path p" "path_image p \<subseteq> s" "x = pathfinish p" | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 628 | shows "homotopic_paths s (p +++ linepath x x) p" | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 629 | using homotopic_paths_rid[of p s] assms by simp | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 630 | |
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 631 | lemma homotopic_paths_lid': | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 632 | "\<lbrakk>path p; path_image p \<subseteq> s; x = pathstart p\<rbrakk> \<Longrightarrow> homotopic_paths s (linepath x x +++ p) p" | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 633 | using homotopic_paths_lid[of p s] by simp | 
| 
35b2143aeec6
An assortment of new material, mostly due to Manuel
 paulson <lp15@cam.ac.uk> parents: 
78516diff
changeset | 634 | |
| 69620 | 635 | proposition homotopic_paths_assoc: | 
| 77684 | 636 | "\<lbrakk>path p; path_image p \<subseteq> S; path q; path_image q \<subseteq> S; path r; path_image r \<subseteq> S; pathfinish p = pathstart q; | 
| 69620 | 637 | pathfinish q = pathstart r\<rbrakk> | 
| 77684 | 638 | \<Longrightarrow> homotopic_paths S (p +++ (q +++ r)) ((p +++ q) +++ r)" | 
| 69620 | 639 | apply (subst homotopic_paths_sym) | 
| 640 | apply (rule homotopic_paths_reparametrize | |
| 71745 | 641 | [where f = "\<lambda>t. if t \<le> 1/2 then inverse 2 *\<^sub>R t | 
| 69620 | 642 | else if t \<le> 3 / 4 then t - (1 / 4) | 
| 643 | else 2 *\<^sub>R t - 1"]) | |
| 72372 | 644 | apply (simp_all del: le_divide_eq_numeral1 add: subset_path_image_join) | 
| 71745 | 645 | apply (rule continuous_on_cases_1 continuous_intros | auto simp: joinpaths_def)+ | 
| 69620 | 646 | done | 
| 647 | ||
| 648 | proposition homotopic_paths_rinv: | |
| 77684 | 649 | assumes "path p" "path_image p \<subseteq> S" | 
| 650 | shows "homotopic_paths S (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" | |
| 69620 | 651 | proof - | 
| 71745 | 652 |   have p: "continuous_on {0..1} p" 
 | 
| 653 | using assms by (auto simp: path_def) | |
| 654 |   let ?A = "{0..1} \<times> {0..1}"
 | |
| 655 | have "continuous_on ?A (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" | |
| 656 | unfolding joinpaths_def subpath_def reversepath_def path_def add_0_right diff_0_right | |
| 657 | proof (rule continuous_on_cases_le) | |
| 658 |     show "continuous_on {x \<in> ?A. snd x \<le> 1/2} (\<lambda>t. p (fst t * (2 * snd t)))"
 | |
| 659 |          "continuous_on {x \<in> ?A. 1/2 \<le> snd x} (\<lambda>t. p (fst t * (1 - (2 * snd t - 1))))"
 | |
| 660 | "continuous_on ?A snd" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 661 | by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+ | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 662 | qed (auto simp: algebra_simps) | 
| 69620 | 663 | then show ?thesis | 
| 664 | using assms | |
| 665 | apply (subst homotopic_paths_sym_eq) | |
| 666 | unfolding homotopic_paths_def homotopic_with_def | |
| 667 | apply (rule_tac x="(\<lambda>y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) | |
| 71745 | 668 | apply (force simp: mult_le_one path_defs joinpaths_def subpath_def reversepath_def) | 
| 69620 | 669 | done | 
| 670 | qed | |
| 671 | ||
| 672 | proposition homotopic_paths_linv: | |
| 77684 | 673 | assumes "path p" "path_image p \<subseteq> S" | 
| 674 | shows "homotopic_paths S (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" | |
| 675 | using homotopic_paths_rinv [of "reversepath p" S] assms by simp | |
| 69620 | 676 | |
| 677 | ||
| 678 | subsection\<open>Homotopy of loops without requiring preservation of endpoints\<close> | |
| 679 | ||
| 70136 | 680 | definition\<^marker>\<open>tag important\<close> homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where | 
| 77684 | 681 | "homotopic_loops S p q \<equiv> | 
| 682 |      homotopic_with_canon (\<lambda>r. pathfinish r = pathstart r) {0..1} S p q"
 | |
| 69620 | 683 | |
| 684 | lemma homotopic_loops: | |
| 77684 | 685 | "homotopic_loops S p q \<longleftrightarrow> | 
| 69620 | 686 |       (\<exists>h. continuous_on ({0..1::real} \<times> {0..1}) h \<and>
 | 
| 77684 | 687 |           image h ({0..1} \<times> {0..1}) \<subseteq> S \<and>
 | 
| 69620 | 688 |           (\<forall>x \<in> {0..1}. h(0,x) = p x) \<and>
 | 
| 689 |           (\<forall>x \<in> {0..1}. h(1,x) = q x) \<and>
 | |
| 690 |           (\<forall>t \<in> {0..1}. pathfinish(h \<circ> Pair t) = pathstart(h \<circ> Pair t)))"
 | |
| 691 | by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) | |
| 692 | ||
| 693 | proposition homotopic_loops_imp_loop: | |
| 77684 | 694 | "homotopic_loops S p q \<Longrightarrow> pathfinish p = pathstart p \<and> pathfinish q = pathstart q" | 
| 69620 | 695 | using homotopic_with_imp_property homotopic_loops_def by blast | 
| 696 | ||
| 697 | proposition homotopic_loops_imp_path: | |
| 77684 | 698 | "homotopic_loops S p q \<Longrightarrow> path p \<and> path q" | 
| 69620 | 699 | unfolding homotopic_loops_def path_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 700 | using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast | 
| 69620 | 701 | |
| 702 | proposition homotopic_loops_imp_subset: | |
| 77684 | 703 | "homotopic_loops S p q \<Longrightarrow> path_image p \<subseteq> S \<and> path_image q \<subseteq> S" | 
| 69620 | 704 | unfolding homotopic_loops_def path_image_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 705 | by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) | 
| 69620 | 706 | |
| 707 | proposition homotopic_loops_refl: | |
| 77684 | 708 | "homotopic_loops S p p \<longleftrightarrow> | 
| 709 | path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 710 | by (simp add: homotopic_loops_def path_image_def path_def) | 
| 69620 | 711 | |
| 77684 | 712 | proposition homotopic_loops_sym: "homotopic_loops S p q \<Longrightarrow> homotopic_loops S q p" | 
| 69620 | 713 | by (simp add: homotopic_loops_def homotopic_with_sym) | 
| 714 | ||
| 77684 | 715 | proposition homotopic_loops_sym_eq: "homotopic_loops S p q \<longleftrightarrow> homotopic_loops S q p" | 
| 69620 | 716 | by (metis homotopic_loops_sym) | 
| 717 | ||
| 718 | proposition homotopic_loops_trans: | |
| 77684 | 719 | "\<lbrakk>homotopic_loops S p q; homotopic_loops S q r\<rbrakk> \<Longrightarrow> homotopic_loops S p r" | 
| 69620 | 720 | unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) | 
| 721 | ||
| 722 | proposition homotopic_loops_subset: | |
| 77684 | 723 | "\<lbrakk>homotopic_loops S p q; S \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t p q" | 
| 78474 | 724 | by (fastforce simp: homotopic_loops) | 
| 69620 | 725 | |
| 726 | proposition homotopic_loops_eq: | |
| 77684 | 727 |    "\<lbrakk>path p; path_image p \<subseteq> S; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk>
 | 
| 728 | \<Longrightarrow> homotopic_loops S p q" | |
| 72372 | 729 | unfolding homotopic_loops_def path_image_def path_def pathstart_def pathfinish_def | 
| 730 | by (auto intro: homotopic_with_eq [OF homotopic_with_refl [where f = p, THEN iffD2]]) | |
| 69620 | 731 | |
| 732 | proposition homotopic_loops_continuous_image: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 733 | "\<lbrakk>homotopic_loops S f g; continuous_on S h; h \<in> S \<rightarrow> t\<rbrakk> \<Longrightarrow> homotopic_loops t (h \<circ> f) (h \<circ> g)" | 
| 69620 | 734 | unfolding homotopic_loops_def | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 735 | by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def image_subset_iff_funcset) | 
| 69620 | 736 | |
| 737 | ||
| 738 | subsection\<open>Relations between the two variants of homotopy\<close> | |
| 739 | ||
| 740 | proposition homotopic_paths_imp_homotopic_loops: | |
| 77684 | 741 | "\<lbrakk>homotopic_paths S p q; pathfinish p = pathstart p; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> homotopic_loops S p q" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 742 | by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def) | 
| 69620 | 743 | |
| 744 | proposition homotopic_loops_imp_homotopic_paths_null: | |
| 77684 | 745 | assumes "homotopic_loops S p (linepath a a)" | 
| 746 | shows "homotopic_paths S p (linepath (pathstart p) (pathstart p))" | |
| 69620 | 747 | proof - | 
| 748 | have "path p" by (metis assms homotopic_loops_imp_path) | |
| 749 | have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) | |
| 77684 | 750 | have pip: "path_image p \<subseteq> S" by (metis assms homotopic_loops_imp_subset) | 
| 71745 | 751 |   let ?A = "{0..1::real} \<times> {0..1::real}"
 | 
| 752 | obtain h where conth: "continuous_on ?A h" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 753 | and hs: "h \<in> ?A \<rightarrow> S" | 
| 77684 | 754 |              and h0[simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(0,x) = p x"
 | 
| 755 |              and h1[simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(1,x) = a"
 | |
| 69620 | 756 |              and ends: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish (h \<circ> Pair t) = pathstart (h \<circ> Pair t)"
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 757 | using assms by (auto simp: homotopic_loops homotopic_with image_subset_iff_funcset) | 
| 69620 | 758 | have conth0: "path (\<lambda>u. h (u, 0))" | 
| 759 | unfolding path_def | |
| 71745 | 760 | proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) | 
| 761 |     show "continuous_on ((\<lambda>x. (x, 0)) ` {0..1}) h"
 | |
| 762 | by (force intro: continuous_on_subset [OF conth]) | |
| 763 | qed (force intro: continuous_intros) | |
| 77684 | 764 | have pih0: "path_image (\<lambda>u. h (u, 0)) \<subseteq> S" | 
| 69620 | 765 | using hs by (force simp: path_image_def) | 
| 71745 | 766 | have c1: "continuous_on ?A (\<lambda>x. h (fst x * snd x, 0))" | 
| 767 | proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) | |
| 768 | show "continuous_on ((\<lambda>x. (fst x * snd x, 0)) ` ?A) h" | |
| 769 | by (force simp: mult_le_one intro: continuous_on_subset [OF conth]) | |
| 770 | qed (force intro: continuous_intros)+ | |
| 771 | have c2: "continuous_on ?A (\<lambda>x. h (fst x - fst x * snd x, 0))" | |
| 772 | proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) | |
| 773 | show "continuous_on ((\<lambda>x. (fst x - fst x * snd x, 0)) ` ?A) h" | |
| 774 | by (auto simp: algebra_simps add_increasing2 mult_left_le intro: continuous_on_subset [OF conth]) | |
| 775 | qed (force intro: continuous_intros) | |
| 69620 | 776 | have [simp]: "\<And>t. \<lbrakk>0 \<le> t \<and> t \<le> 1\<rbrakk> \<Longrightarrow> h (t, 1) = h (t, 0)" | 
| 777 | using ends by (simp add: pathfinish_def pathstart_def) | |
| 778 | have adhoc_le: "c * 4 \<le> 1 + c * (d * 4)" if "\<not> d * 4 \<le> 3" "0 \<le> c" "c \<le> 1" for c d::real | |
| 779 | proof - | |
| 780 | have "c * 3 \<le> c * (d * 4)" using that less_eq_real_def by auto | |
| 781 | with \<open>c \<le> 1\<close> show ?thesis by fastforce | |
| 782 | qed | |
| 71746 | 783 | have *: "\<And>p x. \<lbrakk>path p \<and> path(reversepath p); | 
| 77684 | 784 | path_image p \<subseteq> S \<and> path_image(reversepath p) \<subseteq> S; | 
| 71746 | 785 | pathfinish p = pathstart(linepath a a +++ reversepath p) \<and> | 
| 786 | pathstart(reversepath p) = a \<and> pathstart p = x\<rbrakk> | |
| 77684 | 787 | \<Longrightarrow> homotopic_paths S (p +++ linepath a a +++ reversepath p) (linepath x x)" | 
| 69620 | 788 | by (metis homotopic_paths_lid homotopic_paths_join | 
| 789 | homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) | |
| 77684 | 790 | have 1: "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" | 
| 69620 | 791 | using \<open>path p\<close> homotopic_paths_rid homotopic_paths_sym pip by blast | 
| 77684 | 792 | moreover have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) | 
| 69620 | 793 | (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" | 
| 77684 | 794 | using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" S] | 
| 795 | by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_sym pathstart_join) | |
| 71745 | 796 | moreover | 
| 77684 | 797 | have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) | 
| 69620 | 798 | ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))" | 
| 71745 | 799 | unfolding homotopic_paths_def homotopic_with_def | 
| 800 | proof (intro exI strip conjI) | |
| 71746 | 801 | let ?h = "\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) | 
| 802 | +++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" | |
| 803 | have "continuous_on ?A ?h" | |
| 804 | by (intro continuous_on_homotopic_join_lemma; simp add: path_defs joinpaths_def subpath_def conth c1 c2) | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 805 | moreover have "?h \<in> ?A \<rightarrow> S" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 806 | using hs | 
| 71746 | 807 | unfolding joinpaths_def subpath_def | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 808 | by (force simp: algebra_simps mult_le_one mult_left_le intro: adhoc_le) | 
| 71746 | 809 |   ultimately show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1}))
 | 
| 77684 | 810 | (top_of_set S) ?h" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 811 | by (simp add: subpath_reversepath image_subset_iff_funcset) | 
| 71745 | 812 | qed (use ploop in \<open>simp_all add: reversepath_def path_defs joinpaths_def o_def subpath_def conth c1 c2\<close>) | 
| 77684 | 813 | moreover have "homotopic_paths S ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0))) | 
| 69620 | 814 | (linepath (pathstart p) (pathstart p))" | 
| 77684 | 815 | by (rule *; simp add: pih0 pathstart_def pathfinish_def conth0; simp add: reversepath_def joinpaths_def) | 
| 69620 | 816 | ultimately show ?thesis | 
| 817 | by (blast intro: homotopic_paths_trans) | |
| 818 | qed | |
| 819 | ||
| 820 | proposition homotopic_loops_conjugate: | |
| 77684 | 821 | fixes S :: "'a::real_normed_vector set" | 
| 822 | assumes "path p" "path q" and pip: "path_image p \<subseteq> S" and piq: "path_image q \<subseteq> S" | |
| 71746 | 823 | and pq: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" | 
| 77684 | 824 | shows "homotopic_loops S (p +++ q +++ reversepath p) q" | 
| 69620 | 825 | proof - | 
| 826 |   have contp: "continuous_on {0..1} p"  using \<open>path p\<close> [unfolded path_def] by blast
 | |
| 827 |   have contq: "continuous_on {0..1} q"  using \<open>path q\<close> [unfolded path_def] by blast
 | |
| 71745 | 828 |   let ?A = "{0..1::real} \<times> {0..1::real}"
 | 
| 829 | have c1: "continuous_on ?A (\<lambda>x. p ((1 - fst x) * snd x + fst x))" | |
| 830 | proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) | |
| 831 | show "continuous_on ((\<lambda>x. (1 - fst x) * snd x + fst x) ` ?A) p" | |
| 832 | by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) | |
| 833 | qed (force intro: continuous_intros) | |
| 834 | have c2: "continuous_on ?A (\<lambda>x. p ((fst x - 1) * snd x + 1))" | |
| 835 | proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) | |
| 836 | show "continuous_on ((\<lambda>x. (fst x - 1) * snd x + 1) ` ?A) p" | |
| 837 | by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_left_le_one_le) | |
| 838 | qed (force intro: continuous_intros) | |
| 839 | ||
| 77684 | 840 | have ps1: "\<And>a b. \<lbrakk>b * 2 \<le> 1; 0 \<le> b; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((1 - a) * (2 * b) + a) \<in> S" | 
| 69620 | 841 | using sum_le_prod1 | 
| 842 | by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) | |
| 77684 | 843 | have ps2: "\<And>a b. \<lbrakk>\<not> 4 * b \<le> 3; b \<le> 1; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((a - 1) * (4 * b - 3) + 1) \<in> S" | 
| 69620 | 844 | apply (rule pip [unfolded path_image_def, THEN subsetD]) | 
| 845 | apply (rule image_eqI, blast) | |
| 846 | apply (simp add: algebra_simps) | |
| 77684 | 847 | by (metis add_mono affine_ineq linear mult.commute mult.left_neutral mult_right_mono | 
| 69620 | 848 | add.commute zero_le_numeral) | 
| 77684 | 849 | have qs: "\<And>a b. \<lbrakk>4 * b \<le> 3; \<not> b * 2 \<le> 1\<rbrakk> \<Longrightarrow> q (4 * b - 2) \<in> S" | 
| 69620 | 850 | using path_image_def piq by fastforce | 
| 77684 | 851 | have "homotopic_loops S (p +++ q +++ reversepath p) | 
| 69620 | 852 | (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" | 
| 71746 | 853 | unfolding homotopic_loops_def homotopic_with_def | 
| 854 | proof (intro exI strip conjI) | |
| 855 | let ?h = "(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" | |
| 856 | have "continuous_on ?A (\<lambda>y. q (snd y))" | |
| 857 | by (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) | |
| 858 | then have "continuous_on ?A ?h" | |
| 859 | using pq qloop | |
| 72372 | 860 | by (intro continuous_on_homotopic_join_lemma) (auto simp: path_defs joinpaths_def subpath_def c1 c2) | 
| 77684 | 861 |     then show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h"
 | 
| 71746 | 862 | by (auto simp: joinpaths_def subpath_def ps1 ps2 qs) | 
| 863 | show "?h (1,x) = (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) x" for x | |
| 864 | using pq by (simp add: pathfinish_def subpath_refl) | |
| 865 | qed (auto simp: subpath_reversepath) | |
| 77684 | 866 | moreover have "homotopic_loops S (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" | 
| 69620 | 867 | proof - | 
| 77684 | 868 | have "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q) q" | 
| 69620 | 869 | using \<open>path q\<close> homotopic_paths_lid qloop piq by auto | 
| 77684 | 870 | hence 1: "\<And>f. homotopic_paths S f q \<or> \<not> homotopic_paths S f (linepath (pathfinish q) (pathfinish q) +++ q)" | 
| 69620 | 871 | using homotopic_paths_trans by blast | 
| 77684 | 872 | hence "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" | 
| 873 | by (smt (verit, best) \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_lid | |
| 874 | homotopic_paths_rid homotopic_paths_trans pathstart_join piq qloop) | |
| 69620 | 875 | thus ?thesis | 
| 876 | by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) | |
| 877 | qed | |
| 878 | ultimately show ?thesis | |
| 879 | by (blast intro: homotopic_loops_trans) | |
| 880 | qed | |
| 881 | ||
| 882 | lemma homotopic_paths_loop_parts: | |
| 883 | assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" | |
| 884 | shows "homotopic_paths S p q" | |
| 885 | proof - | |
| 886 | have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" | |
| 887 | using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp | |
| 888 | then have "path p" | |
| 889 | using \<open>path q\<close> homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast | |
| 890 | show ?thesis | |
| 891 | proof (cases "pathfinish p = pathfinish q") | |
| 892 | case True | |
| 77684 | 893 | obtain pipq: "path_image p \<subseteq> S" "path_image q \<subseteq> S" | 
| 69620 | 894 | by (metis Un_subset_iff paths \<open>path p\<close> \<open>path q\<close> homotopic_loops_imp_subset homotopic_paths_imp_path loops | 
| 77684 | 895 | path_image_join path_image_reversepath path_imp_reversepath path_join_eq) | 
| 69620 | 896 | have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" | 
| 897 | using \<open>path p\<close> \<open>path_image p \<subseteq> S\<close> homotopic_paths_rid homotopic_paths_sym by blast | |
| 898 | moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" | |
| 899 | by (simp add: True \<open>path p\<close> \<open>path q\<close> pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) | |
| 900 | moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" | |
| 901 | by (simp add: True \<open>path p\<close> \<open>path q\<close> homotopic_paths_assoc pipq) | |
| 902 | moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" | |
| 903 | by (simp add: \<open>path q\<close> homotopic_paths_join paths pipq) | |
| 904 | ultimately show ?thesis | |
| 77684 | 905 | by (metis \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_lid homotopic_paths_trans path_join_path_ends pathfinish_linepath pipq(2)) | 
| 69620 | 906 | next | 
| 907 | case False | |
| 908 | then show ?thesis | |
| 909 | using \<open>path q\<close> homotopic_loops_imp_path loops path_join_path_ends by fastforce | |
| 910 | qed | |
| 911 | qed | |
| 912 | ||
| 913 | ||
| 70136 | 914 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Homotopy of "nearby" function, paths and loops\<close> | 
| 69620 | 915 | |
| 916 | lemma homotopic_with_linear: | |
| 917 | fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" | |
| 71746 | 918 | assumes contf: "continuous_on S f" | 
| 919 | and contg:"continuous_on S g" | |
| 920 | and sub: "\<And>x. x \<in> S \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t" | |
| 921 | shows "homotopic_with_canon (\<lambda>z. True) S t f g" | |
| 72372 | 922 | unfolding homotopic_with_def | 
| 69620 | 923 | apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) | 
| 72372 | 924 | using sub closed_segment_def | 
| 925 | by (fastforce intro: continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] | |
| 926 | continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]) | |
| 69620 | 927 | |
| 928 | lemma homotopic_paths_linear: | |
| 929 | fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" | |
| 930 | assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" | |
| 71746 | 931 |           "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
 | 
| 932 | shows "homotopic_paths S g h" | |
| 69620 | 933 | using assms | 
| 934 | unfolding path_def | |
| 935 | apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) | |
| 936 | apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R (g \<circ> snd) y + (fst y) *\<^sub>R (h \<circ> snd) y)" in exI) | |
| 937 | apply (intro conjI subsetI continuous_intros; force) | |
| 938 | done | |
| 939 | ||
| 940 | lemma homotopic_loops_linear: | |
| 941 | fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" | |
| 942 | assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" | |
| 71746 | 943 |           "\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S"
 | 
| 944 | shows "homotopic_loops S g h" | |
| 69620 | 945 | using assms | 
| 72372 | 946 | unfolding path_defs homotopic_loops_def homotopic_with_def | 
| 69620 | 947 | apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) | 
| 72372 | 948 | by (force simp: closed_segment_def intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) | 
| 69620 | 949 | |
| 950 | lemma homotopic_paths_nearby_explicit: | |
| 71746 | 951 | assumes \<section>: "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" | 
| 952 |       and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
 | |
| 953 | shows "homotopic_paths S g h" | |
| 77684 | 954 | using homotopic_paths_linear [OF \<section>] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) | 
| 69620 | 955 | |
| 956 | lemma homotopic_loops_nearby_explicit: | |
| 71746 | 957 | assumes \<section>: "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" | 
| 958 |       and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)"
 | |
| 959 | shows "homotopic_loops S g h" | |
| 77684 | 960 | using homotopic_loops_linear [OF \<section>] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) | 
| 69620 | 961 | |
| 962 | lemma homotopic_nearby_paths: | |
| 963 | fixes g h :: "real \<Rightarrow> 'a::euclidean_space" | |
| 71746 | 964 | assumes "path g" "open S" "path_image g \<subseteq> S" | 
| 69620 | 965 | shows "\<exists>e. 0 < e \<and> | 
| 966 | (\<forall>h. path h \<and> | |
| 967 | pathstart h = pathstart g \<and> pathfinish h = pathfinish g \<and> | |
| 71746 | 968 |                     (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths S g h)"
 | 
| 69620 | 969 | proof - | 
| 71746 | 970 | obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y" | 
| 971 | using separate_compact_closed [of "path_image g" "-S"] assms by force | |
| 69620 | 972 | show ?thesis | 
| 71768 | 973 | using e [unfolded dist_norm] \<open>e > 0\<close> | 
| 72372 | 974 | by (fastforce simp: path_image_def intro!: homotopic_paths_nearby_explicit assms exI) | 
| 69620 | 975 | qed | 
| 976 | ||
| 977 | lemma homotopic_nearby_loops: | |
| 978 | fixes g h :: "real \<Rightarrow> 'a::euclidean_space" | |
| 71746 | 979 | assumes "path g" "open S" "path_image g \<subseteq> S" "pathfinish g = pathstart g" | 
| 69620 | 980 | shows "\<exists>e. 0 < e \<and> | 
| 981 | (\<forall>h. path h \<and> pathfinish h = pathstart h \<and> | |
| 71746 | 982 |                     (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops S g h)"
 | 
| 69620 | 983 | proof - | 
| 71746 | 984 | obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y" | 
| 985 | using separate_compact_closed [of "path_image g" "-S"] assms by force | |
| 69620 | 986 | show ?thesis | 
| 71768 | 987 | using e [unfolded dist_norm] \<open>e > 0\<close> | 
| 72372 | 988 | by (fastforce simp: path_image_def intro!: homotopic_loops_nearby_explicit assms exI) | 
| 69620 | 989 | qed | 
| 990 | ||
| 991 | ||
| 992 | subsection\<open> Homotopy and subpaths\<close> | |
| 993 | ||
| 994 | lemma homotopic_join_subpaths1: | |
| 77684 | 995 | assumes "path g" and pag: "path_image g \<subseteq> S" | 
| 69620 | 996 |       and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" "u \<le> v" "v \<le> w"
 | 
| 77684 | 997 | shows "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" | 
| 69620 | 998 | proof - | 
| 999 | have 1: "t * 2 \<le> 1 \<Longrightarrow> u + t * (v * 2) \<le> v + t * (u * 2)" for t | |
| 1000 | using affine_ineq \<open>u \<le> v\<close> by fastforce | |
| 1001 | have 2: "t * 2 > 1 \<Longrightarrow> u + (2*t - 1) * v \<le> v + (2*t - 1) * w" for t | |
| 1002 | by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \<open>u \<le> v\<close> \<open>v \<le> w\<close>) | |
| 1003 | have t2: "\<And>t::real. t*2 = 1 \<Longrightarrow> t = 1/2" by auto | |
| 71746 | 1004 | have "homotopic_paths (path_image g) (subpath u v g +++ subpath v w g) (subpath u w g)" | 
| 1005 | proof (cases "w = u") | |
| 1006 | case True | |
| 1007 | then show ?thesis | |
| 1008 | by (metis \<open>path g\<close> homotopic_paths_rinv path_image_subpath_subset path_subpath pathstart_subpath reversepath_subpath subpath_refl u v) | |
| 1009 | next | |
| 1010 | case False | |
| 1011 | let ?f = "\<lambda>t. if t \<le> 1/2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t | |
| 1012 | else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))" | |
| 1013 | show ?thesis | |
| 1014 | proof (rule homotopic_paths_sym [OF homotopic_paths_reparametrize [where f = ?f]]) | |
| 1015 | show "path (subpath u w g)" | |
| 1016 | using assms(1) path_subpath u w(1) by blast | |
| 1017 | show "path_image (subpath u w g) \<subseteq> path_image g" | |
| 1018 | by (meson path_image_subpath_subset u w(1)) | |
| 1019 |       show "continuous_on {0..1} ?f"
 | |
| 1020 | unfolding split_01 | |
| 1021 | by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def dest!: t2)+ | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1022 |       show "?f \<in> {0..1} \<rightarrow> {0..1}"
 | 
| 71746 | 1023 | using False assms | 
| 1024 | by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2) | |
| 1025 |       show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \<in> {0..1}" for t 
 | |
| 1026 | using assms | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1027 | unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute) | 
| 71746 | 1028 | qed (use False in auto) | 
| 1029 | qed | |
| 1030 | then show ?thesis | |
| 1031 | by (rule homotopic_paths_subset [OF _ pag]) | |
| 69620 | 1032 | qed | 
| 1033 | ||
| 1034 | lemma homotopic_join_subpaths2: | |
| 77684 | 1035 | assumes "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" | 
| 1036 | shows "homotopic_paths S (subpath w v g +++ subpath v u g) (subpath w u g)" | |
| 1037 | by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) | |
| 69620 | 1038 | |
| 1039 | lemma homotopic_join_subpaths3: | |
| 77684 | 1040 | assumes hom: "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" | 
| 1041 | and "path g" and pag: "path_image g \<subseteq> S" | |
| 69620 | 1042 |       and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}"
 | 
| 77684 | 1043 | shows "homotopic_paths S (subpath v w g +++ subpath w u g) (subpath v u g)" | 
| 69620 | 1044 | proof - | 
| 77684 | 1045 | obtain wvg: "path (subpath w v g)" "path_image (subpath w v g) \<subseteq> S" | 
| 1046 | and wug: "path (subpath w u g)" "path_image (subpath w u g) \<subseteq> S" | |
| 1047 | and vug: "path (subpath v u g)" "path_image (subpath v u g) \<subseteq> S" | |
| 1048 | by (meson \<open>path g\<close> pag path_image_subpath_subset path_subpath subset_trans u v w) | |
| 1049 | have "homotopic_paths S (subpath u w g +++ subpath w v g) | |
| 1050 | ((subpath u v g +++ subpath v w g) +++ subpath w v g)" | |
| 1051 | by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg) | |
| 78474 | 1052 | also have "homotopic_paths S \<dots> (subpath u v g +++ subpath v w g +++ subpath w v g)" | 
| 77684 | 1053 | using wvg vug \<open>path g\<close> | 
| 1054 | by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath | |
| 1055 | pathfinish_subpath pathstart_subpath u v w) | |
| 78474 | 1056 | also have "homotopic_paths S \<dots> (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" | 
| 77684 | 1057 | using wvg vug \<open>path g\<close> | 
| 1058 | by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute | |
| 1059 | path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v) | |
| 78474 | 1060 | also have "homotopic_paths S \<dots> (subpath u v g)" | 
| 77684 | 1061 | using vug \<open>path g\<close> by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v) | 
| 1062 | finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" . | |
| 69620 | 1063 | then show ?thesis | 
| 1064 | using homotopic_join_subpaths2 by blast | |
| 1065 | qed | |
| 1066 | ||
| 1067 | proposition homotopic_join_subpaths: | |
| 77684 | 1068 |    "\<lbrakk>path g; path_image g \<subseteq> S; u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
 | 
| 1069 | \<Longrightarrow> homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" | |
| 1070 | by (smt (verit, del_insts) homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3) | |
| 69620 | 1071 | |
| 1072 | text\<open>Relating homotopy of trivial loops to path-connectedness.\<close> | |
| 1073 | ||
| 1074 | lemma path_component_imp_homotopic_points: | |
| 71768 | 1075 | assumes "path_component S a b" | 
| 1076 | shows "homotopic_loops S (linepath a a) (linepath b b)" | |
| 1077 | proof - | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1078 |   obtain g :: "real \<Rightarrow> 'a" where g: "continuous_on {0..1} g" "g \<in> {0..1} \<rightarrow> S" "g 0 = a" "g 1 = b"
 | 
| 71768 | 1079 | using assms by (auto simp: path_defs) | 
| 1080 |   then have "continuous_on ({0..1} \<times> {0..1}) (g \<circ> fst)"
 | |
| 1081 | by (fastforce intro!: continuous_intros)+ | |
| 1082 | with g show ?thesis | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1083 | by (auto simp: homotopic_loops_def homotopic_with_def path_defs Pi_iff) | 
| 71768 | 1084 | qed | 
| 69620 | 1085 | |
| 1086 | lemma homotopic_loops_imp_path_component_value: | |
| 78474 | 1087 | "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> path_component S (p t) (q t)" | 
| 1088 | apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) | |
| 1089 | apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI) | |
| 1090 | apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) | |
| 1091 | done | |
| 69620 | 1092 | |
| 1093 | lemma homotopic_points_eq_path_component: | |
| 71768 | 1094 | "homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow> path_component S a b" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1095 | using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce | 
| 69620 | 1096 | |
| 1097 | lemma path_connected_eq_homotopic_points: | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1098 | "path_connected S \<longleftrightarrow> | 
| 69620 | 1099 | (\<forall>a b. a \<in> S \<and> b \<in> S \<longrightarrow> homotopic_loops S (linepath a a) (linepath b b))" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1100 | by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) | 
| 69620 | 1101 | |
| 1102 | ||
| 1103 | subsection\<open>Simply connected sets\<close> | |
| 1104 | ||
| 70136 | 1105 | text\<^marker>\<open>tag important\<close>\<open>defined as "all loops are homotopic (as loops)\<close> | 
| 1106 | ||
| 1107 | definition\<^marker>\<open>tag important\<close> simply_connected where | |
| 69620 | 1108 | "simply_connected S \<equiv> | 
| 1109 | \<forall>p q. path p \<and> pathfinish p = pathstart p \<and> path_image p \<subseteq> S \<and> | |
| 1110 | path q \<and> pathfinish q = pathstart q \<and> path_image q \<subseteq> S | |
| 1111 | \<longrightarrow> homotopic_loops S p q" | |
| 1112 | ||
| 1113 | lemma simply_connected_empty [iff]: "simply_connected {}"
 | |
| 1114 | by (simp add: simply_connected_def) | |
| 1115 | ||
| 1116 | lemma simply_connected_imp_path_connected: | |
| 1117 | fixes S :: "_::real_normed_vector set" | |
| 1118 | shows "simply_connected S \<Longrightarrow> path_connected S" | |
| 78474 | 1119 | by (simp add: simply_connected_def path_connected_eq_homotopic_points) | 
| 69620 | 1120 | |
| 1121 | lemma simply_connected_imp_connected: | |
| 1122 | fixes S :: "_::real_normed_vector set" | |
| 1123 | shows "simply_connected S \<Longrightarrow> connected S" | |
| 78474 | 1124 | by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) | 
| 69620 | 1125 | |
| 1126 | lemma simply_connected_eq_contractible_loop_any: | |
| 1127 | fixes S :: "_::real_normed_vector set" | |
| 1128 | shows "simply_connected S \<longleftrightarrow> | |
| 71768 | 1129 | (\<forall>p a. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<and> a \<in> S | 
| 69620 | 1130 | \<longrightarrow> homotopic_loops S p (linepath a a))" | 
| 71768 | 1131 | (is "?lhs = ?rhs") | 
| 1132 | proof | |
| 1133 | assume ?rhs then show ?lhs | |
| 1134 | unfolding simply_connected_def | |
| 1135 | by (metis pathfinish_in_path_image subsetD homotopic_loops_trans homotopic_loops_sym) | |
| 78474 | 1136 | qed (force simp: simply_connected_def) | 
| 69620 | 1137 | |
| 1138 | lemma simply_connected_eq_contractible_loop_some: | |
| 1139 | fixes S :: "_::real_normed_vector set" | |
| 1140 | shows "simply_connected S \<longleftrightarrow> | |
| 1141 | path_connected S \<and> | |
| 1142 | (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p | |
| 1143 | \<longrightarrow> (\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)))" | |
| 71746 | 1144 | (is "?lhs = ?rhs") | 
| 1145 | proof | |
| 1146 | assume ?lhs | |
| 1147 | then show ?rhs | |
| 1148 | using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected) | |
| 1149 | next | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1150 | assume ?rhs | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1151 | then show ?lhs | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1152 | by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any) | 
| 71746 | 1153 | qed | 
| 69620 | 1154 | |
| 1155 | lemma simply_connected_eq_contractible_loop_all: | |
| 1156 | fixes S :: "_::real_normed_vector set" | |
| 1157 | shows "simply_connected S \<longleftrightarrow> | |
| 1158 |          S = {} \<or>
 | |
| 1159 | (\<exists>a \<in> S. \<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p | |
| 1160 | \<longrightarrow> homotopic_loops S p (linepath a a))" | |
| 78474 | 1161 | by (meson ex_in_conv homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) | 
| 69620 | 1162 | |
| 1163 | lemma simply_connected_eq_contractible_path: | |
| 1164 | fixes S :: "_::real_normed_vector set" | |
| 1165 | shows "simply_connected S \<longleftrightarrow> | |
| 1166 | path_connected S \<and> | |
| 1167 | (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p | |
| 1168 | \<longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p)))" | |
| 71746 | 1169 | (is "?lhs = ?rhs") | 
| 1170 | proof | |
| 1171 | assume ?lhs | |
| 1172 | then show ?rhs | |
| 1173 | unfolding simply_connected_imp_path_connected | |
| 1174 | by (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) | |
| 1175 | next | |
| 1176 | assume ?rhs | |
| 1177 | then show ?lhs | |
| 1178 | using homotopic_paths_imp_homotopic_loops simply_connected_eq_contractible_loop_some by fastforce | |
| 1179 | qed | |
| 69620 | 1180 | |
| 1181 | lemma simply_connected_eq_homotopic_paths: | |
| 1182 | fixes S :: "_::real_normed_vector set" | |
| 1183 | shows "simply_connected S \<longleftrightarrow> | |
| 1184 | path_connected S \<and> | |
| 1185 | (\<forall>p q. path p \<and> path_image p \<subseteq> S \<and> | |
| 1186 | path q \<and> path_image q \<subseteq> S \<and> | |
| 1187 | pathstart q = pathstart p \<and> pathfinish q = pathfinish p | |
| 1188 | \<longrightarrow> homotopic_paths S p q)" | |
| 1189 | (is "?lhs = ?rhs") | |
| 1190 | proof | |
| 1191 | assume ?lhs | |
| 1192 | then have pc: "path_connected S" | |
| 1193 | and *: "\<And>p. \<lbrakk>path p; path_image p \<subseteq> S; | |
| 1194 | pathfinish p = pathstart p\<rbrakk> | |
| 1195 | \<Longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p))" | |
| 1196 | by (auto simp: simply_connected_eq_contractible_path) | |
| 1197 | have "homotopic_paths S p q" | |
| 1198 | if "path p" "path_image p \<subseteq> S" "path q" | |
| 1199 | "path_image q \<subseteq> S" "pathstart q = pathstart p" | |
| 1200 | "pathfinish q = pathfinish p" for p q | |
| 1201 | proof - | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1202 | have "homotopic_paths S p (p +++ reversepath q +++ q)" | 
| 69620 | 1203 | using that | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1204 | by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1205 | homotopic_paths_trans pathstart_linepath) | 
| 78474 | 1206 | also have "homotopic_paths S \<dots> ((p +++ reversepath q) +++ q)" | 
| 69620 | 1207 | by (simp add: that homotopic_paths_assoc) | 
| 78474 | 1208 | also have "homotopic_paths S \<dots> (linepath (pathstart q) (pathstart q) +++ q)" | 
| 69620 | 1209 | using * [of "p +++ reversepath q"] that | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1210 | by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) | 
| 78474 | 1211 | also have "homotopic_paths S \<dots> q" | 
| 69620 | 1212 | using that homotopic_paths_lid by blast | 
| 1213 | finally show ?thesis . | |
| 1214 | qed | |
| 1215 | then show ?rhs | |
| 1216 | by (blast intro: pc *) | |
| 1217 | next | |
| 1218 | assume ?rhs | |
| 1219 | then show ?lhs | |
| 1220 | by (force simp: simply_connected_eq_contractible_path) | |
| 1221 | qed | |
| 1222 | ||
| 1223 | proposition simply_connected_Times: | |
| 1224 | fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" | |
| 1225 | assumes S: "simply_connected S" and T: "simply_connected T" | |
| 1226 | shows "simply_connected(S \<times> T)" | |
| 1227 | proof - | |
| 1228 | have "homotopic_loops (S \<times> T) p (linepath (a, b) (a, b))" | |
| 1229 | if "path p" "path_image p \<subseteq> S \<times> T" "p 1 = p 0" "a \<in> S" "b \<in> T" | |
| 1230 | for p a b | |
| 1231 | proof - | |
| 1232 | have "path (fst \<circ> p)" | |
| 71746 | 1233 | by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \<open>path p\<close>]) | 
| 69620 | 1234 | moreover have "path_image (fst \<circ> p) \<subseteq> S" | 
| 78474 | 1235 | using that by (force simp: path_image_def) | 
| 69620 | 1236 | ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)" | 
| 1237 | using S that | |
| 71746 | 1238 | by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) | 
| 69620 | 1239 | have "path (snd \<circ> p)" | 
| 71746 | 1240 | by (simp add: continuous_on_snd Path_Connected.path_continuous_image [OF \<open>path p\<close>]) | 
| 69620 | 1241 | moreover have "path_image (snd \<circ> p) \<subseteq> T" | 
| 71768 | 1242 | using that by (force simp: path_image_def) | 
| 69620 | 1243 | ultimately have p2: "homotopic_loops T (snd \<circ> p) (linepath b b)" | 
| 1244 | using T that | |
| 71746 | 1245 | by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) | 
| 69620 | 1246 | show ?thesis | 
| 71746 | 1247 | using p1 p2 unfolding homotopic_loops | 
| 1248 | apply clarify | |
| 72372 | 1249 | subgoal for h k | 
| 1250 | by (rule_tac x="\<lambda>z. (h z, k z)" in exI) (force intro: continuous_intros simp: path_defs) | |
| 69620 | 1251 | done | 
| 1252 | qed | |
| 1253 | with assms show ?thesis | |
| 1254 | by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) | |
| 1255 | qed | |
| 1256 | ||
| 1257 | ||
| 1258 | subsection\<open>Contractible sets\<close> | |
| 1259 | ||
| 70136 | 1260 | definition\<^marker>\<open>tag important\<close> contractible where | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1261 | "contractible S \<equiv> \<exists>a. homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)" | 
| 69620 | 1262 | |
| 1263 | proposition contractible_imp_simply_connected: | |
| 1264 | fixes S :: "_::real_normed_vector set" | |
| 1265 | assumes "contractible S" shows "simply_connected S" | |
| 1266 | proof (cases "S = {}")
 | |
| 1267 | case True then show ?thesis by force | |
| 1268 | next | |
| 1269 | case False | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1270 | obtain a where a: "homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)" | 
| 69620 | 1271 | using assms by (force simp: contractible_def) | 
| 1272 | then have "a \<in> S" | |
| 78336 | 1273 | using False homotopic_with_imp_funspace2 by fastforce | 
| 71746 | 1274 | have "\<forall>p. path p \<and> | 
| 1275 | path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<longrightarrow> | |
| 1276 | homotopic_loops S p (linepath a a)" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1277 | using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) | 
| 69620 | 1278 | apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI) | 
| 71746 | 1279 | apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset) | 
| 69620 | 1280 | done | 
| 71746 | 1281 | with \<open>a \<in> S\<close> show ?thesis | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1282 | by (auto simp: simply_connected_eq_contractible_loop_all False) | 
| 69620 | 1283 | qed | 
| 1284 | ||
| 1285 | corollary contractible_imp_connected: | |
| 1286 | fixes S :: "_::real_normed_vector set" | |
| 1287 | shows "contractible S \<Longrightarrow> connected S" | |
| 78474 | 1288 | by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) | 
| 69620 | 1289 | |
| 1290 | lemma contractible_imp_path_connected: | |
| 1291 | fixes S :: "_::real_normed_vector set" | |
| 1292 | shows "contractible S \<Longrightarrow> path_connected S" | |
| 78474 | 1293 | by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) | 
| 69620 | 1294 | |
| 1295 | lemma nullhomotopic_through_contractible: | |
| 1296 | fixes S :: "_::topological_space set" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1297 | assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1298 | and g: "continuous_on T g" "g \<in> T \<rightarrow> U" | 
| 69620 | 1299 | and T: "contractible T" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1300 | obtains c where "homotopic_with_canon (\<lambda>h. True) S U (g \<circ> f) (\<lambda>x. c)" | 
| 69620 | 1301 | proof - | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1302 | obtain b where b: "homotopic_with_canon (\<lambda>x. True) T T id (\<lambda>x. b)" | 
| 69620 | 1303 | using assms by (force simp: contractible_def) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1304 | have "homotopic_with_canon (\<lambda>f. True) T U (g \<circ> id) (g \<circ> (\<lambda>x. b))" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1305 | by (metis b continuous_map_subtopology_eu g homotopic_with_compose_continuous_map_left image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1306 | then have "homotopic_with_canon (\<lambda>f. True) S U (g \<circ> id \<circ> f) (g \<circ> (\<lambda>x. b) \<circ> f)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1307 | by (simp add: f homotopic_with_compose_continuous_map_right image_subset_iff_funcset) | 
| 69620 | 1308 | then show ?thesis | 
| 1309 | by (simp add: comp_def that) | |
| 1310 | qed | |
| 1311 | ||
| 1312 | lemma nullhomotopic_into_contractible: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1313 | assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" | 
| 69620 | 1314 | and T: "contractible T" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1315 | obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)" | 
| 71746 | 1316 | by (rule nullhomotopic_through_contractible [OF f, of id T]) (use assms in auto) | 
| 69620 | 1317 | |
| 1318 | lemma nullhomotopic_from_contractible: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1319 | assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" | 
| 69620 | 1320 | and S: "contractible S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1321 | obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)" | 
| 71768 | 1322 | by (auto simp: comp_def intro: nullhomotopic_through_contractible [OF continuous_on_id _ f S]) | 
| 69620 | 1323 | |
| 1324 | lemma homotopic_through_contractible: | |
| 1325 | fixes S :: "_::real_normed_vector set" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1326 | assumes "continuous_on S f1" "f1 \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1327 | "continuous_on T g1" "g1 \<in> T \<rightarrow> U" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1328 | "continuous_on S f2" "f2 \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1329 | "continuous_on T g2" "g2 \<in> T \<rightarrow> U" | 
| 69620 | 1330 | "contractible T" "path_connected U" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1331 | shows "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (g2 \<circ> f2)" | 
| 69620 | 1332 | proof - | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1333 | obtain c1 where c1: "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (\<lambda>x. c1)" | 
| 71746 | 1334 | by (rule nullhomotopic_through_contractible [of S f1 T g1 U]) (use assms in auto) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1335 | obtain c2 where c2: "homotopic_with_canon (\<lambda>h. True) S U (g2 \<circ> f2) (\<lambda>x. c2)" | 
| 71746 | 1336 | by (rule nullhomotopic_through_contractible [of S f2 T g2 U]) (use assms in auto) | 
| 1337 |   have "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)"
 | |
| 69620 | 1338 |   proof (cases "S = {}")
 | 
| 1339 | case True then show ?thesis by force | |
| 1340 | next | |
| 1341 | case False | |
| 1342 | with c1 c2 have "c1 \<in> U" "c2 \<in> U" | |
| 78336 | 1343 | using homotopic_with_imp_continuous_maps | 
| 1344 | by (metis PiE equals0I homotopic_with_imp_funspace2)+ | |
| 69620 | 1345 | with \<open>path_connected U\<close> show ?thesis by blast | 
| 1346 | qed | |
| 71746 | 1347 | then have "homotopic_with_canon (\<lambda>h. True) S U (\<lambda>x. c2) (\<lambda>x. c1)" | 
| 78474 | 1348 | by (auto simp: path_component homotopic_constant_maps) | 
| 71746 | 1349 | then show ?thesis | 
| 1350 | using c1 c2 homotopic_with_symD homotopic_with_trans by blast | |
| 69620 | 1351 | qed | 
| 1352 | ||
| 1353 | lemma homotopic_into_contractible: | |
| 1354 | fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1355 | assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1356 | and g: "continuous_on S g" "g \<in> S \<rightarrow> T" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1357 | and T: "contractible T" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1358 | shows "homotopic_with_canon (\<lambda>h. True) S T f g" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1359 | using homotopic_through_contractible [of S f T id T g id] | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1360 | by (simp add: assms contractible_imp_path_connected) | 
| 69620 | 1361 | |
| 1362 | lemma homotopic_from_contractible: | |
| 1363 | fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1364 | assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1365 | and g: "continuous_on S g" "g \<in> S \<rightarrow> T" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1366 | and "contractible S" "path_connected T" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1367 | shows "homotopic_with_canon (\<lambda>h. True) S T f g" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1368 | using homotopic_through_contractible [of S id S f T id g] | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1369 | by (simp add: assms contractible_imp_path_connected) | 
| 69620 | 1370 | |
| 71233 | 1371 | subsection\<open>Starlike sets\<close> | 
| 1372 | ||
| 1373 | definition\<^marker>\<open>tag important\<close> "starlike S \<longleftrightarrow> (\<exists>a\<in>S. \<forall>x\<in>S. closed_segment a x \<subseteq> S)" | |
| 1374 | ||
| 1375 | lemma starlike_UNIV [simp]: "starlike UNIV" | |
| 1376 | by (simp add: starlike_def) | |
| 1377 | ||
| 1378 | lemma convex_imp_starlike: | |
| 1379 |   "convex S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> starlike S"
 | |
| 1380 | unfolding convex_contains_segment starlike_def by auto | |
| 1381 | ||
| 1382 | lemma starlike_convex_tweak_boundary_points: | |
| 1383 | fixes S :: "'a::euclidean_space set" | |
| 1384 |   assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S"
 | |
| 1385 | shows "starlike T" | |
| 1386 | proof - | |
| 1387 |   have "rel_interior S \<noteq> {}"
 | |
| 1388 | by (simp add: assms rel_interior_eq_empty) | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1389 | with ST obtain a where a: "a \<in> rel_interior S" and "a \<in> T" by blast | 
| 71746 | 1390 | have "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" | 
| 1391 | by (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) (use assms in auto) | |
| 1392 | then have "\<forall>x\<in>T. a \<in> T \<and> open_segment a x \<subseteq> T" | |
| 71768 | 1393 | using ST by (blast intro: a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) | 
| 71746 | 1394 | then show ?thesis | 
| 1395 | unfolding starlike_def using bexI [OF _ \<open>a \<in> T\<close>] | |
| 1396 | by (simp add: closed_segment_eq_open) | |
| 71233 | 1397 | qed | 
| 1398 | ||
| 69620 | 1399 | lemma starlike_imp_contractible_gen: | 
| 1400 | fixes S :: "'a::real_normed_vector set" | |
| 1401 | assumes S: "starlike S" | |
| 1402 | and P: "\<And>a T. \<lbrakk>a \<in> S; 0 \<le> T; T \<le> 1\<rbrakk> \<Longrightarrow> P(\<lambda>x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1403 | obtains a where "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)" | 
| 69620 | 1404 | proof - | 
| 1405 | obtain a where "a \<in> S" and a: "\<And>x. x \<in> S \<Longrightarrow> closed_segment a x \<subseteq> S" | |
| 1406 | using S by (auto simp: starlike_def) | |
| 71768 | 1407 | have "\<And>t b. 0 \<le> t \<and> t \<le> 1 \<Longrightarrow> | 
| 1408 | \<exists>u. (1 - t) *\<^sub>R b + t *\<^sub>R a = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" | |
| 1409 | by (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) | |
| 1410 |   then have "(\<lambda>y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \<times> S) \<subseteq> S"
 | |
| 1411 | using a [unfolded closed_segment_def] by force | |
| 71746 | 1412 | then have "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)" | 
| 69620 | 1413 | using \<open>a \<in> S\<close> | 
| 71768 | 1414 | unfolding homotopic_with_def | 
| 69620 | 1415 | apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) | 
| 78474 | 1416 | apply (force simp: P intro: continuous_intros) | 
| 69620 | 1417 | done | 
| 71746 | 1418 | then show ?thesis | 
| 1419 | using that by blast | |
| 69620 | 1420 | qed | 
| 1421 | ||
| 1422 | lemma starlike_imp_contractible: | |
| 1423 | fixes S :: "'a::real_normed_vector set" | |
| 1424 | shows "starlike S \<Longrightarrow> contractible S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1425 | using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) | 
| 69620 | 1426 | |
| 1427 | lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" | |
| 1428 | by (simp add: starlike_imp_contractible) | |
| 1429 | ||
| 1430 | lemma starlike_imp_simply_connected: | |
| 1431 | fixes S :: "'a::real_normed_vector set" | |
| 1432 | shows "starlike S \<Longrightarrow> simply_connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1433 | by (simp add: contractible_imp_simply_connected starlike_imp_contractible) | 
| 69620 | 1434 | |
| 1435 | lemma convex_imp_simply_connected: | |
| 1436 | fixes S :: "'a::real_normed_vector set" | |
| 1437 | shows "convex S \<Longrightarrow> simply_connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1438 | using convex_imp_starlike starlike_imp_simply_connected by blast | 
| 69620 | 1439 | |
| 1440 | lemma starlike_imp_path_connected: | |
| 1441 | fixes S :: "'a::real_normed_vector set" | |
| 1442 | shows "starlike S \<Longrightarrow> path_connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1443 | by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) | 
| 69620 | 1444 | |
| 1445 | lemma starlike_imp_connected: | |
| 1446 | fixes S :: "'a::real_normed_vector set" | |
| 1447 | shows "starlike S \<Longrightarrow> connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1448 | by (simp add: path_connected_imp_connected starlike_imp_path_connected) | 
| 69620 | 1449 | |
| 1450 | lemma is_interval_simply_connected_1: | |
| 1451 | fixes S :: "real set" | |
| 1452 | shows "is_interval S \<longleftrightarrow> simply_connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1453 | by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected) | 
| 69620 | 1454 | |
| 1455 | lemma contractible_empty [simp]: "contractible {}"
 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1456 | by (simp add: contractible_def homotopic_on_emptyI) | 
| 69620 | 1457 | |
| 1458 | lemma contractible_convex_tweak_boundary_points: | |
| 1459 | fixes S :: "'a::euclidean_space set" | |
| 1460 | assumes "convex S" and TS: "rel_interior S \<subseteq> T" "T \<subseteq> closure S" | |
| 1461 | shows "contractible T" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1462 | by (metis assms closure_eq_empty contractible_empty empty_subsetI | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1463 | starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym) | 
| 69620 | 1464 | |
| 1465 | lemma convex_imp_contractible: | |
| 1466 | fixes S :: "'a::real_normed_vector set" | |
| 1467 | shows "convex S \<Longrightarrow> contractible S" | |
| 1468 | using contractible_empty convex_imp_starlike starlike_imp_contractible by blast | |
| 1469 | ||
| 1470 | lemma contractible_sing [simp]: | |
| 1471 | fixes a :: "'a::real_normed_vector" | |
| 1472 |   shows "contractible {a}"
 | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1473 | by (rule convex_imp_contractible [OF convex_singleton]) | 
| 69620 | 1474 | |
| 1475 | lemma is_interval_contractible_1: | |
| 1476 | fixes S :: "real set" | |
| 1477 | shows "is_interval S \<longleftrightarrow> contractible S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1478 | using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1479 | is_interval_simply_connected_1 by auto | 
| 69620 | 1480 | |
| 1481 | lemma contractible_Times: | |
| 1482 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 1483 | assumes S: "contractible S" and T: "contractible T" | |
| 1484 | shows "contractible (S \<times> T)" | |
| 1485 | proof - | |
| 1486 |   obtain a h where conth: "continuous_on ({0..1} \<times> S) h"
 | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1487 |              and hsub: "h \<in> ({0..1} \<times> S) \<rightarrow> S"
 | 
| 69620 | 1488 | and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x" | 
| 1489 | and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (1::real, x) = a" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1490 | using S by (force simp: contractible_def homotopic_with) | 
| 69620 | 1491 |   obtain b k where contk: "continuous_on ({0..1} \<times> T) k"
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1492 |              and ksub: "k \<in> ({0..1} \<times> T) \<rightarrow> T"
 | 
| 69620 | 1493 | and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x" | 
| 1494 | and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (1::real, x) = b" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 1495 | using T by (force simp: contractible_def homotopic_with) | 
| 69620 | 1496 | show ?thesis | 
| 1497 | apply (simp add: contractible_def homotopic_with) | |
| 1498 | apply (rule exI [where x=a]) | |
| 1499 | apply (rule exI [where x=b]) | |
| 1500 | apply (rule exI [where x = "\<lambda>z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) | |
| 1501 | using hsub ksub | |
| 71768 | 1502 | apply (fastforce intro!: continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) | 
| 69620 | 1503 | done | 
| 1504 | qed | |
| 1505 | ||
| 1506 | ||
| 1507 | subsection\<open>Local versions of topological properties in general\<close> | |
| 1508 | ||
| 70136 | 1509 | definition\<^marker>\<open>tag important\<close> locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 69620 | 1510 | where | 
| 1511 | "locally P S \<equiv> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1512 | \<forall>w x. openin (top_of_set S) w \<and> x \<in> w | 
| 78474 | 1513 | \<longrightarrow> (\<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w)" | 
| 69620 | 1514 | |
| 1515 | lemma locallyI: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1516 | assumes "\<And>w x. \<lbrakk>openin (top_of_set S) w; x \<in> w\<rbrakk> | 
| 78474 | 1517 | \<Longrightarrow> \<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w" | 
| 69620 | 1518 | shows "locally P S" | 
| 1519 | using assms by (force simp: locally_def) | |
| 1520 | ||
| 1521 | lemma locallyE: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1522 | assumes "locally P S" "openin (top_of_set S) w" "x \<in> w" | 
| 78474 | 1523 | obtains U V where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> w" | 
| 69620 | 1524 | using assms unfolding locally_def by meson | 
| 1525 | ||
| 1526 | lemma locally_mono: | |
| 71769 | 1527 | assumes "locally P S" "\<And>T. P T \<Longrightarrow> Q T" | 
| 69620 | 1528 | shows "locally Q S" | 
| 1529 | by (metis assms locally_def) | |
| 1530 | ||
| 1531 | lemma locally_open_subset: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1532 | assumes "locally P S" "openin (top_of_set S) t" | 
| 69620 | 1533 | shows "locally P t" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1534 | by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans) | 
| 69620 | 1535 | |
| 1536 | lemma locally_diff_closed: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1537 | "\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> locally P (S - t)" | 
| 69620 | 1538 | using locally_open_subset closedin_def by fastforce | 
| 1539 | ||
| 1540 | lemma locally_empty [iff]: "locally P {}"
 | |
| 1541 | by (simp add: locally_def openin_subtopology) | |
| 1542 | ||
| 1543 | lemma locally_singleton [iff]: | |
| 1544 | fixes a :: "'a::metric_space" | |
| 1545 |   shows "locally P {a} \<longleftrightarrow> P {a}"
 | |
| 71768 | 1546 | proof - | 
| 1547 |   have "\<forall>x::real. \<not> 0 < x \<Longrightarrow> P {a}"
 | |
| 1548 | using zero_less_one by blast | |
| 1549 | then show ?thesis | |
| 1550 | unfolding locally_def | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1551 | by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) | 
| 71768 | 1552 | qed | 
| 69620 | 1553 | |
| 1554 | lemma locally_iff: | |
| 1555 | "locally P S \<longleftrightarrow> | |
| 71746 | 1556 | (\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T)))" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1557 | by (smt (verit) locally_def openin_open) | 
| 69620 | 1558 | |
| 1559 | lemma locally_Int: | |
| 71746 | 1560 | assumes S: "locally P S" and T: "locally P T" | 
| 1561 | and P: "\<And>S T. P S \<and> P T \<Longrightarrow> P(S \<inter> T)" | |
| 1562 | shows "locally P (S \<inter> T)" | |
| 1563 | unfolding locally_iff | |
| 1564 | proof clarify | |
| 1565 | fix A x | |
| 1566 | assume "open A" "x \<in> A" "x \<in> S" "x \<in> T" | |
| 1567 | then obtain U1 V1 U2 V2 | |
| 1568 | where "open U1" "P V1" "x \<in> S \<inter> U1" "S \<inter> U1 \<subseteq> V1 \<and> V1 \<subseteq> S \<inter> A" | |
| 1569 | "open U2" "P V2" "x \<in> T \<inter> U2" "T \<inter> U2 \<subseteq> V2 \<and> V2 \<subseteq> T \<inter> A" | |
| 1570 | using S T unfolding locally_iff by (meson IntI) | |
| 1571 | then have "S \<inter> T \<inter> (U1 \<inter> U2) \<subseteq> V1 \<inter> V2" "V1 \<inter> V2 \<subseteq> S \<inter> T \<inter> A" "x \<in> S \<inter> T \<inter> (U1 \<inter> U2)" | |
| 1572 | by blast+ | |
| 1573 | moreover have "P (V1 \<inter> V2)" | |
| 1574 | by (simp add: P \<open>P V1\<close> \<open>P V2\<close>) | |
| 1575 | ultimately show "\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> T \<inter> U \<and> S \<inter> T \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T \<inter> A)" | |
| 1576 | using \<open>open U1\<close> \<open>open U2\<close> by blast | |
| 1577 | qed | |
| 1578 | ||
| 69620 | 1579 | |
| 1580 | lemma locally_Times: | |
| 1581 |   fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set"
 | |
| 1582 | assumes PS: "locally P S" and QT: "locally Q T" and R: "\<And>S T. P S \<and> Q T \<Longrightarrow> R(S \<times> T)" | |
| 1583 | shows "locally R (S \<times> T)" | |
| 1584 | unfolding locally_def | |
| 1585 | proof (clarify) | |
| 1586 | fix W x y | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1587 | assume W: "openin (top_of_set (S \<times> T)) W" and xy: "(x, y) \<in> W" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1588 | then obtain U V where "openin (top_of_set S) U" "x \<in> U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1589 | "openin (top_of_set T) V" "y \<in> V" "U \<times> V \<subseteq> W" | 
| 69620 | 1590 | using Times_in_interior_subtopology by metis | 
| 1591 | then obtain U1 U2 V1 V2 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1592 | where opeS: "openin (top_of_set S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1593 | and opeT: "openin (top_of_set T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V" | 
| 69620 | 1594 | by (meson PS QT locallyE) | 
| 71768 | 1595 | then have "openin (top_of_set (S \<times> T)) (U1 \<times> V1)" | 
| 1596 | by (simp add: openin_Times) | |
| 1597 | moreover have "R (U2 \<times> V2)" | |
| 1598 | by (simp add: R opeS opeT) | |
| 1599 | moreover have "U1 \<times> V1 \<subseteq> U2 \<times> V2 \<and> U2 \<times> V2 \<subseteq> W" | |
| 1600 | using opeS opeT \<open>U \<times> V \<subseteq> W\<close> by auto | |
| 1601 | ultimately show "\<exists>U V. openin (top_of_set (S \<times> T)) U \<and> R V \<and> (x,y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W" | |
| 1602 | using opeS opeT by auto | |
| 69620 | 1603 | qed | 
| 1604 | ||
| 1605 | ||
| 1606 | proposition homeomorphism_locally_imp: | |
| 71746 | 1607 | fixes S :: "'a::metric_space set" and T :: "'b::t2_space set" | 
| 1608 | assumes S: "locally P S" and hom: "homeomorphism S T f g" | |
| 69620 | 1609 | and Q: "\<And>S S'. \<lbrakk>P S; homeomorphism S S' f g\<rbrakk> \<Longrightarrow> Q S'" | 
| 71746 | 1610 | shows "locally Q T" | 
| 69620 | 1611 | proof (clarsimp simp: locally_def) | 
| 1612 | fix W y | |
| 71746 | 1613 | assume "y \<in> W" and "openin (top_of_set T) W" | 
| 1614 | then obtain A where T: "open A" "W = T \<inter> A" | |
| 69620 | 1615 | by (force simp: openin_open) | 
| 71746 | 1616 | then have "W \<subseteq> T" by auto | 
| 1617 | have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = T" "continuous_on S f" | |
| 1618 | and g: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y" "g ` T = S" "continuous_on T g" | |
| 69620 | 1619 | using hom by (auto simp: homeomorphism_def) | 
| 1620 | have gw: "g ` W = S \<inter> f -` W" | |
| 71746 | 1621 | using \<open>W \<subseteq> T\<close> g by force | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1622 | have "openin (top_of_set S) (g ` W)" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1623 | using \<open>openin (top_of_set T) W\<close> continuous_on_open f gw by auto | 
| 71746 | 1624 | then obtain U V | 
| 1625 | where osu: "openin (top_of_set S) U" and uv: "P V" "g y \<in> U" "U \<subseteq> V" "V \<subseteq> g ` W" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1626 | by (metis S \<open>y \<in> W\<close> image_eqI locallyE) | 
| 71746 | 1627 | have "V \<subseteq> S" using uv by (simp add: gw) | 
| 1628 |   have fv: "f ` V = T \<inter> {x. g x \<in> V}"
 | |
| 1629 | using \<open>f ` S = T\<close> f \<open>V \<subseteq> S\<close> by auto | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1630 | have contvf: "continuous_on V f" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1631 | using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast | 
| 71746 | 1632 | have "openin (top_of_set (g ` T)) U" | 
| 1633 | using \<open>g ` T = S\<close> by (simp add: osu) | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1634 | then have "openin (top_of_set T) (T \<inter> g -` U)" | 
| 71746 | 1635 | using \<open>continuous_on T g\<close> continuous_on_open [THEN iffD1] by blast | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1636 | moreover have "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W" | 
| 71768 | 1637 | proof (intro exI conjI) | 
| 78474 | 1638 | show "f ` V \<subseteq> W" | 
| 1639 | using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto | |
| 1640 | then have contvg: "continuous_on (f ` V) g" | |
| 1641 | using \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast | |
| 1642 | have "V \<subseteq> g ` f ` V" | |
| 1643 | by (metis \<open>V \<subseteq> S\<close> hom homeomorphism_def homeomorphism_of_subsets order_refl) | |
| 1644 | then have homv: "homeomorphism V (f ` V) f g" | |
| 1645 | using \<open>V \<subseteq> S\<close> f by (auto simp: homeomorphism_def contvf contvg) | |
| 71768 | 1646 | show "Q (f ` V)" | 
| 1647 | using Q homv \<open>P V\<close> by blast | |
| 1648 | show "y \<in> T \<inter> g -` U" | |
| 1649 | using T(2) \<open>y \<in> W\<close> \<open>g y \<in> U\<close> by blast | |
| 1650 | show "T \<inter> g -` U \<subseteq> f ` V" | |
| 1651 | using g(1) image_iff uv(3) by fastforce | |
| 1652 | qed | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1653 | ultimately show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1654 | by meson | 
| 69620 | 1655 | qed | 
| 1656 | ||
| 1657 | lemma homeomorphism_locally: | |
| 1658 | fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1659 | assumes "homeomorphism S T f g" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1660 | and "\<And>S T. homeomorphism S T f g \<Longrightarrow> (P S \<longleftrightarrow> Q T)" | 
| 71746 | 1661 | shows "locally P S \<longleftrightarrow> locally Q T" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1662 | by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD) | 
| 69620 | 1663 | |
| 1664 | lemma homeomorphic_locally: | |
| 1665 | fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" | |
| 1666 | assumes hom: "S homeomorphic T" | |
| 1667 | and iff: "\<And>X Y. X homeomorphic Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)" | |
| 1668 | shows "locally P S \<longleftrightarrow> locally Q T" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1669 | by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff) | 
| 69620 | 1670 | |
| 1671 | lemma homeomorphic_local_compactness: | |
| 1672 | fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" | |
| 1673 | shows "S homeomorphic T \<Longrightarrow> locally compact S \<longleftrightarrow> locally compact T" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1674 | by (simp add: homeomorphic_compactness homeomorphic_locally) | 
| 69620 | 1675 | |
| 1676 | lemma locally_translation: | |
| 1677 | fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool" | |
| 71746 | 1678 | shows "(\<And>S. P ((+) a ` S) = P S) \<Longrightarrow> locally P ((+) a ` S) = locally P S" | 
| 71768 | 1679 | using homeomorphism_locally [OF homeomorphism_translation] | 
| 1680 | by (metis (full_types) homeomorphism_image2) | |
| 69620 | 1681 | |
| 1682 | lemma locally_injective_linear_image: | |
| 1683 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 1684 | assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S" | |
| 71746 | 1685 | shows "locally P (f ` S) \<longleftrightarrow> locally Q S" | 
| 78474 | 1686 | by (smt (verit) f homeomorphism_image2 homeomorphism_locally iff linear_homeomorphism_image) | 
| 69620 | 1687 | |
| 1688 | lemma locally_open_map_image: | |
| 1689 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | |
| 1690 | assumes P: "locally P S" | |
| 1691 | and f: "continuous_on S f" | |
| 71746 | 1692 | and oo: "\<And>T. openin (top_of_set S) T \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)" | 
| 1693 | and Q: "\<And>T. \<lbrakk>T \<subseteq> S; P T\<rbrakk> \<Longrightarrow> Q(f ` T)" | |
| 69620 | 1694 | shows "locally Q (f ` S)" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1695 | proof (clarsimp simp: locally_def) | 
| 69620 | 1696 | fix W y | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1697 | assume oiw: "openin (top_of_set (f ` S)) W" and "y \<in> W" | 
| 69620 | 1698 | then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1699 | have oivf: "openin (top_of_set S) (S \<inter> f -` W)" | 
| 69620 | 1700 | by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) | 
| 1701 | then obtain x where "x \<in> S" "f x = y" | |
| 1702 | using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast | |
| 1703 | then obtain U V | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1704 | where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1705 | by (metis IntI P \<open>y \<in> W\<close> locallyE oivf vimageI) | 
| 71746 | 1706 | then have "openin (top_of_set (f ` S)) (f ` U)" | 
| 1707 | by (simp add: oo) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1708 | then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)" | 
| 71746 | 1709 | using Q \<open>P V\<close> \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S \<inter> f -` W\<close> \<open>f x = y\<close> \<open>x \<in> U\<close> by blast | 
| 69620 | 1710 | qed | 
| 1711 | ||
| 1712 | ||
| 1713 | subsection\<open>An induction principle for connected sets\<close> | |
| 1714 | ||
| 1715 | proposition connected_induction: | |
| 1716 | assumes "connected S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1717 | and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" | 
| 69620 | 1718 | and opI: "\<And>a. a \<in> S | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1719 | \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> | 
| 69620 | 1720 | (\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<and> Q x \<longrightarrow> Q y)" | 
| 1721 | and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a" | |
| 1722 | shows "Q b" | |
| 1723 | proof - | |
| 71746 | 1724 |   let ?A = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
 | 
| 1725 |   let ?B = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
 | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1726 |   have "?A \<inter> ?B = {}"
 | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72372diff
changeset | 1727 | by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int) | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1728 | moreover have "S \<subseteq> ?A \<union> ?B" | 
| 71746 | 1729 | by clarsimp (meson opI) | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1730 | moreover have "openin (top_of_set S) ?A" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1731 | by (subst openin_subopen, blast) | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1732 | moreover have "openin (top_of_set S) ?B" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1733 | by (subst openin_subopen, blast) | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1734 |   ultimately have "?A = {} \<or> ?B = {}"
 | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1735 | by (metis (no_types, lifting) \<open>connected S\<close> connected_openin) | 
| 71746 | 1736 | then show ?thesis | 
| 1737 | by clarsimp (meson opI etc) | |
| 69620 | 1738 | qed | 
| 1739 | ||
| 1740 | lemma connected_equivalence_relation_gen: | |
| 1741 | assumes "connected S" | |
| 1742 | and etc: "a \<in> S" "b \<in> S" "P a" "P b" | |
| 1743 | and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1744 | and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" | 
| 69620 | 1745 | and opI: "\<And>a. a \<in> S | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1746 | \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> | 
| 69620 | 1747 | (\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<longrightarrow> R x y)" | 
| 1748 | shows "R a b" | |
| 1749 | proof - | |
| 1750 | have "\<And>a b c. \<lbrakk>a \<in> S; P a; b \<in> S; c \<in> S; P b; P c; R a b\<rbrakk> \<Longrightarrow> R a c" | |
| 1751 | apply (rule connected_induction [OF \<open>connected S\<close> opD], simp_all) | |
| 1752 | by (meson trans opI) | |
| 1753 | then show ?thesis by (metis etc opI) | |
| 1754 | qed | |
| 1755 | ||
| 1756 | lemma connected_induction_simple: | |
| 1757 | assumes "connected S" | |
| 1758 | and etc: "a \<in> S" "b \<in> S" "P a" | |
| 1759 | and opI: "\<And>a. a \<in> S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1760 | \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> | 
| 69620 | 1761 | (\<forall>x \<in> T. \<forall>y \<in> T. P x \<longrightarrow> P y)" | 
| 1762 | shows "P b" | |
| 72372 | 1763 | by (rule connected_induction [OF \<open>connected S\<close> _, where P = "\<lambda>x. True"]) | 
| 1764 | (use opI etc in auto) | |
| 69620 | 1765 | |
| 1766 | lemma connected_equivalence_relation: | |
| 1767 | assumes "connected S" | |
| 1768 | and etc: "a \<in> S" "b \<in> S" | |
| 1769 | and sym: "\<And>x y. \<lbrakk>R x y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> R y x" | |
| 1770 | and trans: "\<And>x y z. \<lbrakk>R x y; R y z; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> R x z" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1771 | and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)" | 
| 69620 | 1772 | shows "R a b" | 
| 1773 | proof - | |
| 1774 | have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c" | |
| 78474 | 1775 | by (smt (verit, ccfv_threshold) connected_induction_simple [OF \<open>connected S\<close>] | 
| 1776 | assms openin_imp_subset subset_eq) | |
| 69620 | 1777 | then show ?thesis by (metis etc opI) | 
| 1778 | qed | |
| 1779 | ||
| 1780 | lemma locally_constant_imp_constant: | |
| 1781 | assumes "connected S" | |
| 1782 | and opI: "\<And>a. a \<in> S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1783 | \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)" | 
| 69620 | 1784 | shows "f constant_on S" | 
| 1785 | proof - | |
| 1786 | have "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> f x = f y" | |
| 1787 | apply (rule connected_equivalence_relation [OF \<open>connected S\<close>], simp_all) | |
| 1788 | by (metis opI) | |
| 1789 | then show ?thesis | |
| 1790 | by (metis constant_on_def) | |
| 1791 | qed | |
| 1792 | ||
| 1793 | lemma locally_constant: | |
| 71768 | 1794 | assumes "connected S" | 
| 1795 | shows "locally (\<lambda>U. f constant_on U) S \<longleftrightarrow> f constant_on S" (is "?lhs = ?rhs") | |
| 1796 | proof | |
| 1797 | assume ?lhs | |
| 1798 | then show ?rhs | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1799 | by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff) | 
| 71768 | 1800 | next | 
| 1801 | assume ?rhs then show ?lhs | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1802 | by (metis constant_on_subset locallyI openin_imp_subset order_refl) | 
| 71768 | 1803 | qed | 
| 69620 | 1804 | |
| 1805 | ||
| 1806 | subsection\<open>Basic properties of local compactness\<close> | |
| 1807 | ||
| 1808 | proposition locally_compact: | |
| 77684 | 1809 | fixes S :: "'a :: metric_space set" | 
| 69620 | 1810 | shows | 
| 77684 | 1811 | "locally compact S \<longleftrightarrow> | 
| 1812 | (\<forall>x \<in> S. \<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> | |
| 1813 | openin (top_of_set S) u \<and> compact v)" | |
| 69620 | 1814 | (is "?lhs = ?rhs") | 
| 1815 | proof | |
| 1816 | assume ?lhs | |
| 1817 | then show ?rhs | |
| 71768 | 1818 | by (meson locallyE openin_subtopology_self) | 
| 69620 | 1819 | next | 
| 1820 | assume r [rule_format]: ?rhs | |
| 1821 | have *: "\<exists>u v. | |
| 77684 | 1822 | openin (top_of_set S) u \<and> | 
| 1823 | compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<inter> T" | |
| 1824 | if "open T" "x \<in> S" "x \<in> T" for x T | |
| 69620 | 1825 | proof - | 
| 78474 | 1826 | obtain U V where uv: "x \<in> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" "openin (top_of_set S) U" | 
| 77684 | 1827 | using r [OF \<open>x \<in> S\<close>] by auto | 
| 69620 | 1828 | obtain e where "e>0" and e: "cball x e \<subseteq> T" | 
| 1829 | using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast | |
| 1830 | show ?thesis | |
| 78474 | 1831 | apply (rule_tac x="(S \<inter> ball x e) \<inter> U" in exI) | 
| 1832 | apply (rule_tac x="cball x e \<inter> V" in exI) | |
| 69620 | 1833 | using that \<open>e > 0\<close> e uv | 
| 1834 | apply auto | |
| 1835 | done | |
| 1836 | qed | |
| 1837 | show ?lhs | |
| 71768 | 1838 | by (rule locallyI) (metis "*" Int_iff openin_open) | 
| 69620 | 1839 | qed | 
| 1840 | ||
| 1841 | lemma locally_compactE: | |
| 71746 | 1842 | fixes S :: "'a :: metric_space set" | 
| 1843 | assumes "locally compact S" | |
| 1844 | obtains u v where "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and> | |
| 1845 | openin (top_of_set S) (u x) \<and> compact (v x)" | |
| 1846 | using assms unfolding locally_compact by metis | |
| 69620 | 1847 | |
| 1848 | lemma locally_compact_alt: | |
| 71746 | 1849 | fixes S :: "'a :: heine_borel set" | 
| 1850 | shows "locally compact S \<longleftrightarrow> | |
| 71768 | 1851 | (\<forall>x \<in> S. \<exists>U. x \<in> U \<and> | 
| 1852 | openin (top_of_set S) U \<and> compact(closure U) \<and> closure U \<subseteq> S)" | |
| 78474 | 1853 | by (smt (verit, ccfv_threshold) bounded_subset closure_closed closure_mono closure_subset | 
| 1854 | compact_closure compact_imp_closed order.trans locally_compact) | |
| 69620 | 1855 | |
| 1856 | lemma locally_compact_Int_cball: | |
| 71746 | 1857 | fixes S :: "'a :: heine_borel set" | 
| 1858 | shows "locally compact S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>e. 0 < e \<and> closed(cball x e \<inter> S))" | |
| 69620 | 1859 | (is "?lhs = ?rhs") | 
| 1860 | proof | |
| 72372 | 1861 | assume L: ?lhs | 
| 1862 | then have "\<And>x U V e. \<lbrakk>U \<subseteq> V; V \<subseteq> S; compact V; 0 < e; cball x e \<inter> S \<subseteq> U\<rbrakk> | |
| 1863 | \<Longrightarrow> closed (cball x e \<inter> S)" | |
| 1864 | by (metis compact_Int compact_cball compact_imp_closed inf.absorb_iff2 inf.assoc inf.orderE) | |
| 1865 | with L show ?rhs | |
| 1866 | by (meson locally_compactE openin_contains_cball) | |
| 69620 | 1867 | next | 
| 72372 | 1868 | assume R: ?rhs | 
| 1869 | show ?lhs unfolding locally_compact | |
| 1870 | proof | |
| 1871 | fix x | |
| 1872 | assume "x \<in> S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1873 | then obtain e where "e>0" and "compact (cball x e \<inter> S)" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1874 | by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R) | 
| 72372 | 1875 | moreover have "\<forall>y\<in>ball x e \<inter> S. \<exists>\<epsilon>>0. cball y \<epsilon> \<inter> S \<subseteq> ball x e" | 
| 1876 | by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq) | |
| 1877 | moreover have "openin (top_of_set S) (ball x e \<inter> S)" | |
| 1878 | by (simp add: inf_commute openin_open_Int) | |
| 1879 | ultimately show "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" | |
| 1880 | by (metis Int_iff \<open>0 < e\<close> \<open>x \<in> S\<close> ball_subset_cball centre_in_ball inf_commute inf_le1 inf_mono order_refl) | |
| 1881 | qed | |
| 69620 | 1882 | qed | 
| 1883 | ||
| 1884 | lemma locally_compact_compact: | |
| 71746 | 1885 | fixes S :: "'a :: heine_borel set" | 
| 1886 | shows "locally compact S \<longleftrightarrow> | |
| 72372 | 1887 | (\<forall>K. K \<subseteq> S \<and> compact K | 
| 1888 | \<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> | |
| 1889 | openin (top_of_set S) U \<and> compact V))" | |
| 69620 | 1890 | (is "?lhs = ?rhs") | 
| 1891 | proof | |
| 1892 | assume ?lhs | |
| 1893 | then obtain u v where | |
| 71746 | 1894 | uv: "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and> | 
| 1895 | openin (top_of_set S) (u x) \<and> compact (v x)" | |
| 69620 | 1896 | by (metis locally_compactE) | 
| 72372 | 1897 | have *: "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" | 
| 1898 | if "K \<subseteq> S" "compact K" for K | |
| 69620 | 1899 | proof - | 
| 72372 | 1900 | have "\<And>C. (\<forall>c\<in>C. openin (top_of_set K) c) \<and> K \<subseteq> \<Union>C \<Longrightarrow> | 
| 1901 | \<exists>D\<subseteq>C. finite D \<and> K \<subseteq> \<Union>D" | |
| 69620 | 1902 | using that by (simp add: compact_eq_openin_cover) | 
| 72372 | 1903 | moreover have "\<forall>c \<in> (\<lambda>x. K \<inter> u x) ` K. openin (top_of_set K) c" | 
| 69620 | 1904 | using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) | 
| 72372 | 1905 | moreover have "K \<subseteq> \<Union>((\<lambda>x. K \<inter> u x) ` K)" | 
| 69620 | 1906 | using that by clarsimp (meson subsetCE uv) | 
| 72372 | 1907 | ultimately obtain D where "D \<subseteq> (\<lambda>x. K \<inter> u x) ` K" "finite D" "K \<subseteq> \<Union>D" | 
| 69620 | 1908 | by metis | 
| 72372 | 1909 | then obtain T where T: "T \<subseteq> K" "finite T" "K \<subseteq> \<Union>((\<lambda>x. K \<inter> u x) ` T)" | 
| 69620 | 1910 | by (metis finite_subset_image) | 
| 1911 | have Tuv: "\<Union>(u ` T) \<subseteq> \<Union>(v ` T)" | |
| 72372 | 1912 | using T that by (force dest!: uv) | 
| 1913 | moreover | |
| 1914 | have "openin (top_of_set S) (\<Union> (u ` T))" | |
| 1915 | using T that uv by fastforce | |
| 1916 | moreover | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1917 | obtain "compact (\<Union> (v ` T))" "\<Union> (v ` T) \<subseteq> S" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1918 | by (metis T UN_subset_iff \<open>K \<subseteq> S\<close> compact_UN subset_iff uv) | 
| 72372 | 1919 | ultimately show ?thesis | 
| 1920 | using T by auto | |
| 69620 | 1921 | qed | 
| 1922 | show ?rhs | |
| 1923 | by (blast intro: *) | |
| 1924 | next | |
| 1925 | assume ?rhs | |
| 1926 | then show ?lhs | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1927 | apply (clarsimp simp: locally_compact) | 
| 69620 | 1928 |     apply (drule_tac x="{x}" in spec, simp)
 | 
| 1929 | done | |
| 1930 | qed | |
| 1931 | ||
| 1932 | lemma open_imp_locally_compact: | |
| 71746 | 1933 | fixes S :: "'a :: heine_borel set" | 
| 1934 | assumes "open S" | |
| 1935 | shows "locally compact S" | |
| 69620 | 1936 | proof - | 
| 71746 | 1937 | have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" | 
| 1938 | if "x \<in> S" for x | |
| 69620 | 1939 | proof - | 
| 71746 | 1940 | obtain e where "e>0" and e: "cball x e \<subseteq> S" | 
| 1941 | using open_contains_cball assms \<open>x \<in> S\<close> by blast | |
| 1942 | have ope: "openin (top_of_set S) (ball x e)" | |
| 69620 | 1943 | by (meson e open_ball ball_subset_cball dual_order.trans open_subset) | 
| 1944 | show ?thesis | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1945 | by (meson \<open>0 < e\<close> ball_subset_cball centre_in_ball compact_cball e ope) | 
| 69620 | 1946 | qed | 
| 1947 | show ?thesis | |
| 71746 | 1948 | unfolding locally_compact by (blast intro: *) | 
| 69620 | 1949 | qed | 
| 1950 | ||
| 1951 | lemma closed_imp_locally_compact: | |
| 71746 | 1952 | fixes S :: "'a :: heine_borel set" | 
| 1953 | assumes "closed S" | |
| 1954 | shows "locally compact S" | |
| 69620 | 1955 | proof - | 
| 71746 | 1956 | have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" | 
| 1957 | if "x \<in> S" for x | |
| 72372 | 1958 | apply (rule_tac x = "S \<inter> ball x 1" in exI, rule_tac x = "S \<inter> cball x 1" in exI) | 
| 1959 | using \<open>x \<in> S\<close> assms by auto | |
| 69620 | 1960 | show ?thesis | 
| 72372 | 1961 | unfolding locally_compact by (blast intro: *) | 
| 69620 | 1962 | qed | 
| 1963 | ||
| 1964 | lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" | |
| 1965 | by (simp add: closed_imp_locally_compact) | |
| 1966 | ||
| 1967 | lemma locally_compact_Int: | |
| 71746 | 1968 | fixes S :: "'a :: t2_space set" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1969 | shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<inter> T)" | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1970 | by (simp add: compact_Int locally_Int) | 
| 69620 | 1971 | |
| 1972 | lemma locally_compact_closedin: | |
| 71746 | 1973 | fixes S :: "'a :: heine_borel set" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1974 | shows "\<lbrakk>closedin (top_of_set S) T; locally compact S\<rbrakk> | 
| 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 1975 | \<Longrightarrow> locally compact T" | 
| 71746 | 1976 | unfolding closedin_closed | 
| 1977 | using closed_imp_locally_compact locally_compact_Int by blast | |
| 69620 | 1978 | |
| 1979 | lemma locally_compact_delete: | |
| 71746 | 1980 | fixes S :: "'a :: t1_space set" | 
| 1981 |      shows "locally compact S \<Longrightarrow> locally compact (S - {a})"
 | |
| 69620 | 1982 | by (auto simp: openin_delete locally_open_subset) | 
| 1983 | ||
| 1984 | lemma locally_closed: | |
| 71746 | 1985 | fixes S :: "'a :: heine_borel set" | 
| 1986 | shows "locally closed S \<longleftrightarrow> locally compact S" | |
| 69620 | 1987 | (is "?lhs = ?rhs") | 
| 1988 | proof | |
| 1989 | assume ?lhs | |
| 1990 | then show ?rhs | |
| 72372 | 1991 | unfolding locally_def | 
| 1992 | apply (elim all_forward imp_forward asm_rl exE) | |
| 78474 | 1993 | apply (rename_tac U V) | 
| 1994 | apply (rule_tac x = "U \<inter> ball x 1" in exI) | |
| 1995 | apply (rule_tac x = "V \<inter> cball x 1" in exI) | |
| 69620 | 1996 | apply (force intro: openin_trans) | 
| 1997 | done | |
| 1998 | next | |
| 1999 | assume ?rhs then show ?lhs | |
| 2000 | using compact_eq_bounded_closed locally_mono by blast | |
| 2001 | qed | |
| 2002 | ||
| 2003 | lemma locally_compact_openin_Un: | |
| 2004 | fixes S :: "'a::euclidean_space set" | |
| 78474 | 2005 | assumes LCS: "locally compact S" and LCT: "locally compact T" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2006 | and opS: "openin (top_of_set (S \<union> T)) S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2007 | and opT: "openin (top_of_set (S \<union> T)) T" | 
| 69620 | 2008 | shows "locally compact (S \<union> T)" | 
| 2009 | proof - | |
| 2010 | have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" for x | |
| 2011 | proof - | |
| 2012 | obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" | |
| 2013 | using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast | |
| 2014 | moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> S" | |
| 2015 | by (meson \<open>x \<in> S\<close> opS openin_contains_cball) | |
| 2016 | then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> S" | |
| 2017 | by force | |
| 71768 | 2018 | ultimately have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" | 
| 2019 | by (metis (no_types, lifting) cball_min_Int closed_Int closed_cball inf_assoc inf_commute) | |
| 2020 | then show ?thesis | |
| 2021 | by (metis \<open>0 < e1\<close> \<open>0 < e2\<close> min_def) | |
| 69620 | 2022 | qed | 
| 2023 | moreover have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> T" for x | |
| 2024 | proof - | |
| 2025 | obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" | |
| 2026 | using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast | |
| 2027 | moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> T" | |
| 2028 | by (meson \<open>x \<in> T\<close> opT openin_contains_cball) | |
| 2029 | then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> T" | |
| 2030 | by force | |
| 72372 | 2031 | moreover have "closed (cball x e1 \<inter> (cball x e2 \<inter> T))" | 
| 2032 | by (metis closed_Int closed_cball e1 inf_left_commute) | |
| 69620 | 2033 | ultimately show ?thesis | 
| 72372 | 2034 | by (rule_tac x="min e1 e2" in exI) (simp add: \<open>0 < e2\<close> cball_min_Int inf_assoc) | 
| 69620 | 2035 | qed | 
| 2036 | ultimately show ?thesis | |
| 2037 | by (force simp: locally_compact_Int_cball) | |
| 2038 | qed | |
| 2039 | ||
| 2040 | lemma locally_compact_closedin_Un: | |
| 2041 | fixes S :: "'a::euclidean_space set" | |
| 2042 | assumes LCS: "locally compact S" and LCT:"locally compact T" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2043 | and clS: "closedin (top_of_set (S \<union> T)) S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2044 | and clT: "closedin (top_of_set (S \<union> T)) T" | 
| 69620 | 2045 | shows "locally compact (S \<union> T)" | 
| 2046 | proof - | |
| 2047 | have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x | |
| 2048 | proof - | |
| 2049 | obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" | |
| 2050 | using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast | |
| 2051 | moreover | |
| 2052 | obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)" | |
| 2053 | using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast | |
| 72372 | 2054 | moreover have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 2055 | by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute) | 
| 69620 | 2056 | ultimately show ?thesis | 
| 72372 | 2057 | by (rule_tac x="min e1 e2" in exI) linarith | 
| 69620 | 2058 | qed | 
| 2059 | moreover | |
| 2060 | have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<in> S" "x \<notin> T" for x | |
| 2061 | proof - | |
| 2062 | obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" | |
| 2063 | using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast | |
| 2064 | moreover | |
| 2065 | obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S - T" | |
| 2066 | using clT x by (fastforce simp: openin_contains_cball closedin_def) | |
| 2067 | then have "closed (cball x e2 \<inter> T)" | |
| 2068 | proof - | |
| 2069 |       have "{} = T - (T - cball x e2)"
 | |
| 2070 | using Diff_subset Int_Diff \<open>cball x e2 \<inter> (S \<union> T) \<subseteq> S - T\<close> by auto | |
| 2071 | then show ?thesis | |
| 2072 | by (simp add: Diff_Diff_Int inf_commute) | |
| 2073 | qed | |
| 72372 | 2074 | with e1 have "closed ((cball x e1 \<inter> cball x e2) \<inter> (S \<union> T))" | 
| 2075 | apply (simp add: inf_commute inf_sup_distrib2) | |
| 2076 | by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) | |
| 2077 | then have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" | |
| 2078 | by (simp add: cball_min_Int inf_commute) | |
| 69620 | 2079 | ultimately show ?thesis | 
| 72372 | 2080 | using \<open>0 < e2\<close> by (rule_tac x="min e1 e2" in exI) linarith | 
| 69620 | 2081 | qed | 
| 2082 | moreover | |
| 2083 | have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<notin> S" "x \<in> T" for x | |
| 2084 | proof - | |
| 2085 | obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" | |
| 2086 | using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast | |
| 2087 | moreover | |
| 2088 | obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S \<union> T - S" | |
| 2089 | using clS x by (fastforce simp: openin_contains_cball closedin_def) | |
| 2090 | then have "closed (cball x e2 \<inter> S)" | |
| 2091 | by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) | |
| 72372 | 2092 | with e1 have "closed ((cball x e1 \<inter> cball x e2) \<inter> (S \<union> T))" | 
| 2093 | apply (simp add: inf_commute inf_sup_distrib2) | |
| 2094 | by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) | |
| 2095 | then have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" | |
| 2096 | by (auto simp: cball_min_Int) | |
| 69620 | 2097 | ultimately show ?thesis | 
| 72372 | 2098 | using \<open>0 < e2\<close> by (rule_tac x="min e1 e2" in exI) linarith | 
| 69620 | 2099 | qed | 
| 2100 | ultimately show ?thesis | |
| 2101 | by (auto simp: locally_compact_Int_cball) | |
| 2102 | qed | |
| 2103 | ||
| 2104 | lemma locally_compact_Times: | |
| 2105 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 2106 | shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<times> T)" | |
| 2107 | by (auto simp: compact_Times locally_Times) | |
| 2108 | ||
| 2109 | lemma locally_compact_compact_subopen: | |
| 2110 | fixes S :: "'a :: heine_borel set" | |
| 2111 | shows | |
| 2112 | "locally compact S \<longleftrightarrow> | |
| 2113 | (\<forall>K T. K \<subseteq> S \<and> compact K \<and> open T \<and> K \<subseteq> T | |
| 2114 | \<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2115 | openin (top_of_set S) U \<and> compact V))" | 
| 69620 | 2116 | (is "?lhs = ?rhs") | 
| 2117 | proof | |
| 2118 | assume L: ?lhs | |
| 2119 | show ?rhs | |
| 2120 | proof clarify | |
| 2121 | fix K :: "'a set" and T :: "'a set" | |
| 2122 | assume "K \<subseteq> S" and "compact K" and "open T" and "K \<subseteq> T" | |
| 2123 | obtain U V where "K \<subseteq> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2124 | and ope: "openin (top_of_set S) U" | 
| 69620 | 2125 | using L unfolding locally_compact_compact by (meson \<open>K \<subseteq> S\<close> \<open>compact K\<close>) | 
| 2126 | show "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2127 | openin (top_of_set S) U \<and> compact V" | 
| 69620 | 2128 | proof (intro exI conjI) | 
| 2129 | show "K \<subseteq> U \<inter> T" | |
| 2130 | by (simp add: \<open>K \<subseteq> T\<close> \<open>K \<subseteq> U\<close>) | |
| 2131 | show "U \<inter> T \<subseteq> closure(U \<inter> T)" | |
| 2132 | by (rule closure_subset) | |
| 2133 | show "closure (U \<inter> T) \<subseteq> S" | |
| 2134 | by (metis \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S\<close> \<open>compact V\<close> closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2135 | show "openin (top_of_set S) (U \<inter> T)" | 
| 69620 | 2136 | by (simp add: \<open>open T\<close> ope openin_Int_open) | 
| 2137 | show "compact (closure (U \<inter> T))" | |
| 2138 | by (meson Int_lower1 \<open>U \<subseteq> V\<close> \<open>compact V\<close> bounded_subset compact_closure compact_eq_bounded_closed) | |
| 2139 | qed auto | |
| 2140 | qed | |
| 2141 | next | |
| 2142 | assume ?rhs then show ?lhs | |
| 2143 | unfolding locally_compact_compact | |
| 2144 | by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) | |
| 2145 | qed | |
| 2146 | ||
| 2147 | ||
| 2148 | subsection\<open>Sura-Bura's results about compact components of sets\<close> | |
| 2149 | ||
| 2150 | proposition Sura_Bura_compact: | |
| 2151 | fixes S :: "'a::euclidean_space set" | |
| 2152 | assumes "compact S" and C: "C \<in> components S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2153 |   shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set S) T \<and>
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2154 | closedin (top_of_set S) T}" | 
| 69620 | 2155 | (is "C = \<Inter>?\<T>") | 
| 2156 | proof | |
| 2157 | obtain x where x: "C = connected_component_set S x" and "x \<in> S" | |
| 2158 | using C by (auto simp: components_def) | |
| 2159 | have "C \<subseteq> S" | |
| 2160 | by (simp add: C in_components_subset) | |
| 2161 | have "\<Inter>?\<T> \<subseteq> connected_component_set S x" | |
| 2162 | proof (rule connected_component_maximal) | |
| 2163 | have "x \<in> C" | |
| 2164 | by (simp add: \<open>x \<in> S\<close> x) | |
| 2165 | then show "x \<in> \<Inter>?\<T>" | |
| 2166 | by blast | |
| 2167 | have clo: "closed (\<Inter>?\<T>)" | |
| 2168 | by (simp add: \<open>compact S\<close> closed_Inter closedin_compact_eq compact_imp_closed) | |
| 2169 | have False | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2170 | if K1: "closedin (top_of_set (\<Inter>?\<T>)) K1" and | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2171 | K2: "closedin (top_of_set (\<Inter>?\<T>)) K2" and | 
| 69620 | 2172 |          K12_Int: "K1 \<inter> K2 = {}" and K12_Un: "K1 \<union> K2 = \<Inter>?\<T>" and "K1 \<noteq> {}" "K2 \<noteq> {}"
 | 
| 2173 | for K1 K2 | |
| 2174 | proof - | |
| 2175 | have "closed K1" "closed K2" | |
| 2176 | using closedin_closed_trans clo K1 K2 by blast+ | |
| 2177 |       then obtain V1 V2 where "open V1" "open V2" "K1 \<subseteq> V1" "K2 \<subseteq> V2" and V12: "V1 \<inter> V2 = {}"
 | |
| 2178 |         using separation_normal \<open>K1 \<inter> K2 = {}\<close> by metis
 | |
| 2179 |       have SV12_ne: "(S - (V1 \<union> V2)) \<inter> (\<Inter>?\<T>) \<noteq> {}"
 | |
| 2180 | proof (rule compact_imp_fip) | |
| 2181 | show "compact (S - (V1 \<union> V2))" | |
| 2182 | by (simp add: \<open>open V1\<close> \<open>open V2\<close> \<open>compact S\<close> compact_diff open_Un) | |
| 2183 | show clo\<T>: "closed T" if "T \<in> ?\<T>" for T | |
| 2184 | using that \<open>compact S\<close> | |
| 2185 | by (force intro: closedin_closed_trans simp add: compact_imp_closed) | |
| 2186 |         show "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> \<noteq> {}" if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F>
 | |
| 2187 | proof | |
| 2188 |           assume djo: "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> = {}"
 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2189 | obtain D where opeD: "openin (top_of_set S) D" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2190 | and cloD: "closedin (top_of_set S) D" | 
| 69620 | 2191 | and "C \<subseteq> D" and DV12: "D \<subseteq> V1 \<union> V2" | 
| 2192 |           proof (cases "\<F> = {}")
 | |
| 2193 | case True | |
| 2194 | with \<open>C \<subseteq> S\<close> djo that show ?thesis | |
| 2195 | by force | |
| 2196 | next | |
| 2197 | case False show ?thesis | |
| 2198 | proof | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2199 | show ope: "openin (top_of_set S) (\<Inter>\<F>)" | 
| 69620 | 2200 | using openin_Inter \<open>finite \<F>\<close> False \<F> by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2201 | then show "closedin (top_of_set S) (\<Inter>\<F>)" | 
| 69620 | 2202 | by (meson clo\<T> \<F> closed_Inter closed_subset openin_imp_subset subset_eq) | 
| 2203 | show "C \<subseteq> \<Inter>\<F>" | |
| 2204 | using \<F> by auto | |
| 2205 | show "\<Inter>\<F> \<subseteq> V1 \<union> V2" | |
| 2206 | using ope djo openin_imp_subset by fastforce | |
| 2207 | qed | |
| 2208 | qed | |
| 2209 | have "connected C" | |
| 2210 | by (simp add: x) | |
| 2211 | have "closed D" | |
| 2212 | using \<open>compact S\<close> cloD closedin_closed_trans compact_imp_closed by blast | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2213 | have cloV1: "closedin (top_of_set D) (D \<inter> closure V1)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2214 | and cloV2: "closedin (top_of_set D) (D \<inter> closure V2)" | 
| 69620 | 2215 | by (simp_all add: closedin_closed_Int) | 
| 2216 | moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2" | |
| 2217 | using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12 | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 2218 | by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2219 | ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2220 | and cloDV2: "closedin (top_of_set D) (D \<inter> V2)" | 
| 69620 | 2221 | by metis+ | 
| 2222 | then obtain U1 U2 where "closed U1" "closed U2" | |
| 2223 | and D1: "D \<inter> V1 = D \<inter> U1" and D2: "D \<inter> V2 = D \<inter> U2" | |
| 2224 | by (auto simp: closedin_closed) | |
| 2225 |           have "D \<inter> U1 \<inter> C \<noteq> {}"
 | |
| 2226 | proof | |
| 2227 |             assume "D \<inter> U1 \<inter> C = {}"
 | |
| 2228 | then have *: "C \<subseteq> D \<inter> V2" | |
| 2229 | using D1 DV12 \<open>C \<subseteq> D\<close> by auto | |
| 71746 | 2230 | have 1: "openin (top_of_set S) (D \<inter> V2)" | 
| 2231 | by (simp add: \<open>open V2\<close> opeD openin_Int_open) | |
| 2232 | have 2: "closedin (top_of_set S) (D \<inter> V2)" | |
| 2233 | using cloD cloDV2 closedin_trans by blast | |
| 2234 | have "\<Inter> ?\<T> \<subseteq> D \<inter> V2" | |
| 2235 | by (rule Inter_lower) (use * 1 2 in simp) | |
| 69620 | 2236 | then show False | 
| 2237 |               using K1 V12 \<open>K1 \<noteq> {}\<close> \<open>K1 \<subseteq> V1\<close> closedin_imp_subset by blast
 | |
| 2238 | qed | |
| 2239 |           moreover have "D \<inter> U2 \<inter> C \<noteq> {}"
 | |
| 2240 | proof | |
| 2241 |             assume "D \<inter> U2 \<inter> C = {}"
 | |
| 2242 | then have *: "C \<subseteq> D \<inter> V1" | |
| 2243 | using D2 DV12 \<open>C \<subseteq> D\<close> by auto | |
| 71746 | 2244 | have 1: "openin (top_of_set S) (D \<inter> V1)" | 
| 2245 | by (simp add: \<open>open V1\<close> opeD openin_Int_open) | |
| 2246 | have 2: "closedin (top_of_set S) (D \<inter> V1)" | |
| 2247 | using cloD cloDV1 closedin_trans by blast | |
| 69620 | 2248 | have "\<Inter>?\<T> \<subseteq> D \<inter> V1" | 
| 71746 | 2249 | by (rule Inter_lower) (use * 1 2 in simp) | 
| 69620 | 2250 | then show False | 
| 2251 |               using K2 V12 \<open>K2 \<noteq> {}\<close> \<open>K2 \<subseteq> V2\<close> closedin_imp_subset by blast
 | |
| 2252 | qed | |
| 2253 | ultimately show False | |
| 71746 | 2254 | using \<open>connected C\<close> [unfolded connected_closed, simplified, rule_format, of concl: "D \<inter> U1" "D \<inter> U2"] | 
| 69620 | 2255 | using \<open>C \<subseteq> D\<close> D1 D2 V12 DV12 \<open>closed U1\<close> \<open>closed U2\<close> \<open>closed D\<close> | 
| 2256 | by blast | |
| 2257 | qed | |
| 2258 | qed | |
| 2259 | show False | |
| 2260 | by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \<open>K1 \<subseteq> V1\<close> \<open>K2 \<subseteq> V2\<close> disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) | |
| 2261 | qed | |
| 2262 | then show "connected (\<Inter>?\<T>)" | |
| 2263 | by (auto simp: connected_closedin_eq) | |
| 2264 | show "\<Inter>?\<T> \<subseteq> S" | |
| 2265 | by (fastforce simp: C in_components_subset) | |
| 2266 | qed | |
| 2267 | with x show "\<Inter>?\<T> \<subseteq> C" by simp | |
| 2268 | qed auto | |
| 2269 | ||
| 2270 | ||
| 2271 | corollary Sura_Bura_clopen_subset: | |
| 2272 | fixes S :: "'a::euclidean_space set" | |
| 2273 | assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" | |
| 2274 | and U: "open U" "C \<subseteq> U" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2275 | obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" | 
| 69620 | 2276 | proof (rule ccontr) | 
| 2277 | assume "\<not> thesis" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2278 | with that have neg: "\<nexists>K. openin (top_of_set S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U" | 
| 69620 | 2279 | by metis | 
| 2280 | obtain V K where "C \<subseteq> V" "V \<subseteq> U" "V \<subseteq> K" "K \<subseteq> S" "compact K" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2281 | and opeSV: "openin (top_of_set S) V" | 
| 72372 | 2282 | using S U \<open>compact C\<close> by (meson C in_components_subset locally_compact_compact_subopen) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2283 |   let ?\<T> = "{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> compact T \<and> T \<subseteq> K}"
 | 
| 69620 | 2284 | have CK: "C \<in> components K" | 
| 2285 | by (meson C \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> components_intermediate_subset subset_trans) | |
| 2286 | with \<open>compact K\<close> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2287 |   have "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> closedin (top_of_set K) T}"
 | 
| 69620 | 2288 | by (simp add: Sura_Bura_compact) | 
| 2289 | then have Ceq: "C = \<Inter>?\<T>" | |
| 2290 | by (simp add: closedin_compact_eq \<open>compact K\<close>) | |
| 2291 | obtain W where "open W" and W: "V = S \<inter> W" | |
| 2292 | using opeSV by (auto simp: openin_open) | |
| 2293 |   have "-(U \<inter> W) \<inter> \<Inter>?\<T> \<noteq> {}"
 | |
| 2294 | proof (rule closed_imp_fip_compact) | |
| 2295 |     show "- (U \<inter> W) \<inter> \<Inter>\<F> \<noteq> {}"
 | |
| 2296 | if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F> | |
| 2297 |     proof (cases "\<F> = {}")
 | |
| 2298 | case True | |
| 2299 | have False if "U = UNIV" "W = UNIV" | |
| 2300 | proof - | |
| 2301 | have "V = S" | |
| 2302 | by (simp add: W \<open>W = UNIV\<close>) | |
| 2303 | with neg show False | |
| 2304 | using \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> \<open>V \<subseteq> U\<close> \<open>compact K\<close> by auto | |
| 2305 | qed | |
| 2306 | with True show ?thesis | |
| 2307 | by auto | |
| 2308 | next | |
| 2309 | case False | |
| 2310 | show ?thesis | |
| 2311 | proof | |
| 2312 |         assume "- (U \<inter> W) \<inter> \<Inter>\<F> = {}"
 | |
| 2313 | then have FUW: "\<Inter>\<F> \<subseteq> U \<inter> W" | |
| 2314 | by blast | |
| 2315 | have "C \<subseteq> \<Inter>\<F>" | |
| 2316 | using \<F> by auto | |
| 2317 | moreover have "compact (\<Inter>\<F>)" | |
| 2318 | by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \<F>) | |
| 2319 | moreover have "\<Inter>\<F> \<subseteq> K" | |
| 2320 | using False that(2) by fastforce | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2321 | moreover have opeKF: "openin (top_of_set K) (\<Inter>\<F>)" | 
| 69620 | 2322 | using False \<F> \<open>finite \<F>\<close> by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2323 | then have opeVF: "openin (top_of_set V) (\<Inter>\<F>)" | 
| 69620 | 2324 | using W \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> opeKF \<open>\<Inter>\<F> \<subseteq> K\<close> FUW openin_subset_trans by fastforce | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2325 | then have "openin (top_of_set S) (\<Inter>\<F>)" | 
| 69620 | 2326 | by (metis opeSV openin_trans) | 
| 2327 | moreover have "\<Inter>\<F> \<subseteq> U" | |
| 2328 | by (meson \<open>V \<subseteq> U\<close> opeVF dual_order.trans openin_imp_subset) | |
| 2329 | ultimately show False | |
| 2330 | using neg by blast | |
| 2331 | qed | |
| 2332 | qed | |
| 2333 | qed (use \<open>open W\<close> \<open>open U\<close> in auto) | |
| 2334 | with W Ceq \<open>C \<subseteq> V\<close> \<open>C \<subseteq> U\<close> show False | |
| 2335 | by auto | |
| 2336 | qed | |
| 2337 | ||
| 2338 | ||
| 2339 | corollary Sura_Bura_clopen_subset_alt: | |
| 2340 | fixes S :: "'a::euclidean_space set" | |
| 2341 | assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2342 | and opeSU: "openin (top_of_set S) U" and "C \<subseteq> U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2343 | obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" | 
| 69620 | 2344 | proof - | 
| 2345 | obtain V where "open V" "U = S \<inter> V" | |
| 2346 | using opeSU by (auto simp: openin_open) | |
| 2347 | with \<open>C \<subseteq> U\<close> have "C \<subseteq> V" | |
| 2348 | by auto | |
| 2349 | then show ?thesis | |
| 2350 | using Sura_Bura_clopen_subset [OF S C \<open>compact C\<close> \<open>open V\<close>] | |
| 2351 | by (metis \<open>U = S \<inter> V\<close> inf.bounded_iff openin_imp_subset that) | |
| 2352 | qed | |
| 2353 | ||
| 2354 | corollary Sura_Bura: | |
| 2355 | fixes S :: "'a::euclidean_space set" | |
| 2356 | assumes "locally compact S" "C \<in> components S" "compact C" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2357 |   shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (top_of_set S) K}"
 | 
| 69620 | 2358 | (is "C = ?rhs") | 
| 2359 | proof | |
| 2360 | show "?rhs \<subseteq> C" | |
| 2361 | proof (clarsimp, rule ccontr) | |
| 2362 | fix x | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2363 | assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (top_of_set S) X \<longrightarrow> x \<in> X" | 
| 69620 | 2364 | and "x \<notin> C" | 
| 2365 |     obtain U V where "open U" "open V" "{x} \<subseteq> U" "C \<subseteq> V" "U \<inter> V = {}"
 | |
| 2366 |       using separation_normal [of "{x}" C]
 | |
| 2367 | by (metis Int_empty_left \<open>x \<notin> C\<close> \<open>compact C\<close> closed_empty closed_insert compact_imp_closed insert_disjoint(1)) | |
| 2368 | have "x \<notin> V" | |
| 2369 |       using \<open>U \<inter> V = {}\<close> \<open>{x} \<subseteq> U\<close> by blast
 | |
| 2370 | then show False | |
| 2371 | by (meson "*" Sura_Bura_clopen_subset \<open>C \<subseteq> V\<close> \<open>open V\<close> assms(1) assms(2) assms(3) subsetCE) | |
| 2372 | qed | |
| 2373 | qed blast | |
| 2374 | ||
| 2375 | ||
| 2376 | subsection\<open>Special cases of local connectedness and path connectedness\<close> | |
| 2377 | ||
| 2378 | lemma locally_connected_1: | |
| 2379 | assumes | |
| 72372 | 2380 | "\<And>V x. \<lbrakk>openin (top_of_set S) V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. openin (top_of_set S) U \<and> connected U \<and> x \<in> U \<and> U \<subseteq> V" | 
| 69620 | 2381 | shows "locally connected S" | 
| 72372 | 2382 | by (metis assms locally_def) | 
| 69620 | 2383 | |
| 2384 | lemma locally_connected_2: | |
| 2385 | assumes "locally connected S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2386 | "openin (top_of_set S) t" | 
| 69620 | 2387 | "x \<in> t" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2388 | shows "openin (top_of_set S) (connected_component_set t x)" | 
| 69620 | 2389 | proof - | 
| 2390 |   { fix y :: 'a
 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2391 | let ?SS = "top_of_set S" | 
| 69620 | 2392 | assume 1: "openin ?SS t" | 
| 2393 | "\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" | |
| 2394 | and "connected_component t x y" | |
| 2395 | then have "y \<in> t" and y: "y \<in> connected_component_set t x" | |
| 2396 | using connected_component_subset by blast+ | |
| 2397 | obtain F where | |
| 2398 | "\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" | |
| 2399 | by moura | |
| 2400 | then obtain G where | |
| 2401 | "\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" | |
| 2402 | by moura | |
| 2403 | then have *: "openin ?SS (F y t) \<and> connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" | |
| 2404 | using 1 \<open>y \<in> t\<close> by presburger | |
| 2405 | have "G y t \<subseteq> connected_component_set t y" | |
| 2406 | by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) | |
| 2407 | then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> connected_component_set t x" | |
| 2408 | by (metis (no_types) * connected_component_eq dual_order.trans y) | |
| 2409 | } | |
| 2410 | then show ?thesis | |
| 2411 | using assms openin_subopen by (force simp: locally_def) | |
| 2412 | qed | |
| 2413 | ||
| 2414 | lemma locally_connected_3: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2415 | assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2416 | \<Longrightarrow> openin (top_of_set S) | 
| 69620 | 2417 | (connected_component_set t x)" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2418 | "openin (top_of_set S) v" "x \<in> v" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2419 | shows "\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v" | 
| 69620 | 2420 | using assms connected_component_subset by fastforce | 
| 2421 | ||
| 2422 | lemma locally_connected: | |
| 2423 | "locally connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2424 | (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2425 | \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))" | 
| 69620 | 2426 | by (metis locally_connected_1 locally_connected_2 locally_connected_3) | 
| 2427 | ||
| 2428 | lemma locally_connected_open_connected_component: | |
| 2429 | "locally connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2430 | (\<forall>t x. openin (top_of_set S) t \<and> x \<in> t | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2431 | \<longrightarrow> openin (top_of_set S) (connected_component_set t x))" | 
| 69620 | 2432 | by (metis locally_connected_1 locally_connected_2 locally_connected_3) | 
| 2433 | ||
| 2434 | lemma locally_path_connected_1: | |
| 2435 | assumes | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2436 | "\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2437 | \<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" | 
| 69620 | 2438 | shows "locally path_connected S" | 
| 78474 | 2439 | by (force simp: locally_def dest: assms) | 
| 69620 | 2440 | |
| 2441 | lemma locally_path_connected_2: | |
| 2442 | assumes "locally path_connected S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2443 | "openin (top_of_set S) t" | 
| 69620 | 2444 | "x \<in> t" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2445 | shows "openin (top_of_set S) (path_component_set t x)" | 
| 69620 | 2446 | proof - | 
| 2447 |   { fix y :: 'a
 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2448 | let ?SS = "top_of_set S" | 
| 69620 | 2449 | assume 1: "openin ?SS t" | 
| 2450 | "\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. path_connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" | |
| 2451 | and "path_component t x y" | |
| 2452 | then have "y \<in> t" and y: "y \<in> path_component_set t x" | |
| 2453 | using path_component_mem(2) by blast+ | |
| 2454 | obtain F where | |
| 2455 | "\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. path_connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. path_connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" | |
| 2456 | by moura | |
| 2457 | then obtain G where | |
| 2458 | "\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. path_connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> path_connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" | |
| 2459 | by moura | |
| 2460 | then have *: "openin ?SS (F y t) \<and> path_connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" | |
| 2461 | using 1 \<open>y \<in> t\<close> by presburger | |
| 2462 | have "G y t \<subseteq> path_component_set t y" | |
| 69712 | 2463 | using * path_component_maximal rev_subsetD by blast | 
| 69620 | 2464 | then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> path_component_set t x" | 
| 2465 | by (metis "*" \<open>G y t \<subseteq> path_component_set t y\<close> dual_order.trans path_component_eq y) | |
| 2466 | } | |
| 2467 | then show ?thesis | |
| 2468 | using assms openin_subopen by (force simp: locally_def) | |
| 2469 | qed | |
| 2470 | ||
| 2471 | lemma locally_path_connected_3: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2472 | assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2473 | \<Longrightarrow> openin (top_of_set S) (path_component_set t x)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2474 | "openin (top_of_set S) v" "x \<in> v" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2475 | shows "\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" | 
| 69620 | 2476 | proof - | 
| 2477 | have "path_component v x x" | |
| 2478 | by (meson assms(3) path_component_refl) | |
| 2479 | then show ?thesis | |
| 71746 | 2480 | by (metis assms mem_Collect_eq path_component_subset path_connected_path_component) | 
| 69620 | 2481 | qed | 
| 2482 | ||
| 2483 | proposition locally_path_connected: | |
| 2484 | "locally path_connected S \<longleftrightarrow> | |
| 71769 | 2485 | (\<forall>V x. openin (top_of_set S) V \<and> x \<in> V | 
| 2486 | \<longrightarrow> (\<exists>U. openin (top_of_set S) U \<and> path_connected U \<and> x \<in> U \<and> U \<subseteq> V))" | |
| 69620 | 2487 | by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) | 
| 2488 | ||
| 2489 | proposition locally_path_connected_open_path_component: | |
| 2490 | "locally path_connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2491 | (\<forall>t x. openin (top_of_set S) t \<and> x \<in> t | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2492 | \<longrightarrow> openin (top_of_set S) (path_component_set t x))" | 
| 69620 | 2493 | by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) | 
| 2494 | ||
| 2495 | lemma locally_connected_open_component: | |
| 2496 | "locally connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2497 | (\<forall>t c. openin (top_of_set S) t \<and> c \<in> components t | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2498 | \<longrightarrow> openin (top_of_set S) c)" | 
| 69620 | 2499 | by (metis components_iff locally_connected_open_connected_component) | 
| 2500 | ||
| 2501 | proposition locally_connected_im_kleinen: | |
| 2502 | "locally connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2503 | (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2504 | \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> | 
| 69620 | 2505 | x \<in> u \<and> u \<subseteq> v \<and> | 
| 2506 | (\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> v \<and> x \<in> c \<and> y \<in> c))))" | |
| 2507 | (is "?lhs = ?rhs") | |
| 2508 | proof | |
| 2509 | assume ?lhs | |
| 2510 | then show ?rhs | |
| 78474 | 2511 | by (fastforce simp: locally_connected) | 
| 69620 | 2512 | next | 
| 2513 | assume ?rhs | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2514 | have *: "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> c" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2515 | if "openin (top_of_set S) t" and c: "c \<in> components t" and "x \<in> c" for t c x | 
| 69620 | 2516 | proof - | 
| 2517 | from that \<open>?rhs\<close> [rule_format, of t x] | |
| 2518 | obtain u where u: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2519 | "openin (top_of_set S) u \<and> x \<in> u \<and> u \<subseteq> t \<and> | 
| 69620 | 2520 | (\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> t \<and> x \<in> c \<and> y \<in> c))" | 
| 2521 | using in_components_subset by auto | |
| 2522 | obtain F :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where | |
| 2523 | "\<forall>x y. (\<exists>z. z \<in> x \<and> y = connected_component_set x z) = (F x y \<in> x \<and> y = connected_component_set x (F x y))" | |
| 2524 | by moura | |
| 2525 | then have F: "F t c \<in> t \<and> c = connected_component_set t (F t c)" | |
| 2526 | by (meson components_iff c) | |
| 2527 | obtain G :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where | |
| 2528 | G: "\<forall>x y. (\<exists>z. z \<in> y \<and> z \<notin> x) = (G x y \<in> y \<and> G x y \<notin> x)" | |
| 2529 | by moura | |
| 2530 | have "G c u \<notin> u \<or> G c u \<in> c" | |
| 2531 | using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) | |
| 2532 | then show ?thesis | |
| 2533 | using G u by auto | |
| 2534 | qed | |
| 2535 | show ?lhs | |
| 71769 | 2536 | unfolding locally_connected_open_component by (meson "*" openin_subopen) | 
| 69620 | 2537 | qed | 
| 2538 | ||
| 2539 | proposition locally_path_connected_im_kleinen: | |
| 2540 | "locally path_connected S \<longleftrightarrow> | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2541 | (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2542 | \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> | 
| 69620 | 2543 | x \<in> u \<and> u \<subseteq> v \<and> | 
| 2544 | (\<forall>y. y \<in> u \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> v \<and> | |
| 2545 | pathstart p = x \<and> pathfinish p = y))))" | |
| 2546 | (is "?lhs = ?rhs") | |
| 2547 | proof | |
| 2548 | assume ?lhs | |
| 2549 | then show ?rhs | |
| 2550 | apply (simp add: locally_path_connected path_connected_def) | |
| 2551 | apply (erule all_forward ex_forward imp_forward conjE | simp)+ | |
| 2552 | by (meson dual_order.trans) | |
| 2553 | next | |
| 2554 | assume ?rhs | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2555 | have *: "\<exists>T. openin (top_of_set S) T \<and> | 
| 69620 | 2556 | x \<in> T \<and> T \<subseteq> path_component_set u z" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2557 | if "openin (top_of_set S) u" and "z \<in> u" and c: "path_component u z x" for u z x | 
| 69620 | 2558 | proof - | 
| 2559 | have "x \<in> u" | |
| 2560 | by (meson c path_component_mem(2)) | |
| 2561 | with that \<open>?rhs\<close> [rule_format, of u x] | |
| 2562 | obtain U where U: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2563 | "openin (top_of_set S) U \<and> x \<in> U \<and> U \<subseteq> u \<and> | 
| 69620 | 2564 | (\<forall>y. y \<in> U \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> u \<and> pathstart p = x \<and> pathfinish p = y))" | 
| 2565 | by blast | |
| 2566 | show ?thesis | |
| 71769 | 2567 | by (metis U c mem_Collect_eq path_component_def path_component_eq subsetI) | 
| 69620 | 2568 | qed | 
| 2569 | show ?lhs | |
| 71769 | 2570 | unfolding locally_path_connected_open_path_component | 
| 2571 | using "*" openin_subopen by fastforce | |
| 69620 | 2572 | qed | 
| 2573 | ||
| 2574 | lemma locally_path_connected_imp_locally_connected: | |
| 2575 | "locally path_connected S \<Longrightarrow> locally connected S" | |
| 2576 | using locally_mono path_connected_imp_connected by blast | |
| 2577 | ||
| 2578 | lemma locally_connected_components: | |
| 2579 | "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally connected c" | |
| 2580 | by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) | |
| 2581 | ||
| 2582 | lemma locally_path_connected_components: | |
| 2583 | "\<lbrakk>locally path_connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally path_connected c" | |
| 2584 | by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) | |
| 2585 | ||
| 2586 | lemma locally_path_connected_connected_component: | |
| 2587 | "locally path_connected S \<Longrightarrow> locally path_connected (connected_component_set S x)" | |
| 2588 | by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) | |
| 2589 | ||
| 2590 | lemma open_imp_locally_path_connected: | |
| 2591 | fixes S :: "'a :: real_normed_vector set" | |
| 71769 | 2592 | assumes "open S" | 
| 2593 | shows "locally path_connected S" | |
| 2594 | proof (rule locally_mono) | |
| 2595 | show "locally convex S" | |
| 2596 | using assms unfolding locally_def | |
| 2597 | by (meson open_ball centre_in_ball convex_ball openE open_subset openin_imp_subset openin_open_trans subset_trans) | |
| 2598 | show "\<And>T::'a set. convex T \<Longrightarrow> path_connected T" | |
| 2599 | using convex_imp_path_connected by blast | |
| 2600 | qed | |
| 69620 | 2601 | |
| 2602 | lemma open_imp_locally_connected: | |
| 2603 | fixes S :: "'a :: real_normed_vector set" | |
| 2604 | shows "open S \<Longrightarrow> locally connected S" | |
| 2605 | by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) | |
| 2606 | ||
| 2607 | lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" | |
| 2608 | by (simp add: open_imp_locally_path_connected) | |
| 2609 | ||
| 2610 | lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" | |
| 2611 | by (simp add: open_imp_locally_connected) | |
| 2612 | ||
| 2613 | lemma openin_connected_component_locally_connected: | |
| 2614 | "locally connected S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2615 | \<Longrightarrow> openin (top_of_set S) (connected_component_set S x)" | 
| 71769 | 2616 | by (metis connected_component_eq_empty locally_connected_2 openin_empty openin_subtopology_self) | 
| 69620 | 2617 | |
| 2618 | lemma openin_components_locally_connected: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2619 | "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> openin (top_of_set S) c" | 
| 69620 | 2620 | using locally_connected_open_component openin_subtopology_self by blast | 
| 2621 | ||
| 2622 | lemma openin_path_component_locally_path_connected: | |
| 2623 | "locally path_connected S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2624 | \<Longrightarrow> openin (top_of_set S) (path_component_set S x)" | 
| 69620 | 2625 | by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) | 
| 2626 | ||
| 2627 | lemma closedin_path_component_locally_path_connected: | |
| 71769 | 2628 | assumes "locally path_connected S" | 
| 2629 | shows "closedin (top_of_set S) (path_component_set S x)" | |
| 2630 | proof - | |
| 2631 |   have "openin (top_of_set S) (\<Union> ({path_component_set S y |y. y \<in> S} - {path_component_set S x}))"
 | |
| 2632 | using locally_path_connected_2 assms by fastforce | |
| 2633 | then show ?thesis | |
| 2634 | by (simp add: closedin_def path_component_subset complement_path_component_Union) | |
| 2635 | qed | |
| 69620 | 2636 | |
| 2637 | lemma convex_imp_locally_path_connected: | |
| 2638 | fixes S :: "'a:: real_normed_vector set" | |
| 71769 | 2639 | assumes "convex S" | 
| 2640 | shows "locally path_connected S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 2641 | proof (clarsimp simp: locally_path_connected) | 
| 71769 | 2642 | fix V x | 
| 2643 | assume "openin (top_of_set S) V" and "x \<in> V" | |
| 2644 | then obtain T e where "V = S \<inter> T" "x \<in> S" "0 < e" "ball x e \<subseteq> T" | |
| 2645 | by (metis Int_iff openE openin_open) | |
| 2646 | then have "openin (top_of_set S) (S \<inter> ball x e)" "path_connected (S \<inter> ball x e)" | |
| 2647 | by (simp_all add: assms convex_Int convex_imp_path_connected openin_open_Int) | |
| 2648 | then show "\<exists>U. openin (top_of_set S) U \<and> path_connected U \<and> x \<in> U \<and> U \<subseteq> V" | |
| 2649 | using \<open>0 < e\<close> \<open>V = S \<inter> T\<close> \<open>ball x e \<subseteq> T\<close> \<open>x \<in> S\<close> by auto | |
| 2650 | qed | |
| 69620 | 2651 | |
| 2652 | lemma convex_imp_locally_connected: | |
| 2653 | fixes S :: "'a:: real_normed_vector set" | |
| 2654 | shows "convex S \<Longrightarrow> locally connected S" | |
| 2655 | by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) | |
| 2656 | ||
| 2657 | ||
| 2658 | subsection\<open>Relations between components and path components\<close> | |
| 2659 | ||
| 2660 | lemma path_component_eq_connected_component: | |
| 2661 | assumes "locally path_connected S" | |
| 2662 | shows "(path_component S x = connected_component S x)" | |
| 2663 | proof (cases "x \<in> S") | |
| 2664 | case True | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2665 | have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" | 
| 71769 | 2666 | proof (rule openin_subset_trans) | 
| 2667 | show "openin (top_of_set S) (path_component_set S x)" | |
| 2668 | by (simp add: True assms locally_path_connected_2) | |
| 2669 | show "connected_component_set S x \<subseteq> S" | |
| 2670 | by (simp add: connected_component_subset) | |
| 2671 | qed (simp add: path_component_subset_connected_component) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2672 | moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" | 
| 71769 | 2673 | proof (rule closedin_subset_trans [of S]) | 
| 2674 | show "closedin (top_of_set S) (path_component_set S x)" | |
| 2675 | by (simp add: assms closedin_path_component_locally_path_connected) | |
| 2676 | show "connected_component_set S x \<subseteq> S" | |
| 2677 | by (simp add: connected_component_subset) | |
| 2678 | qed (simp add: path_component_subset_connected_component) | |
| 69620 | 2679 | ultimately have *: "path_component_set S x = connected_component_set S x" | 
| 2680 | by (metis connected_connected_component connected_clopen True path_component_eq_empty) | |
| 2681 | then show ?thesis | |
| 2682 | by blast | |
| 2683 | next | |
| 2684 | case False then show ?thesis | |
| 2685 | by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) | |
| 2686 | qed | |
| 2687 | ||
| 2688 | lemma path_component_eq_connected_component_set: | |
| 2689 | "locally path_connected S \<Longrightarrow> (path_component_set S x = connected_component_set S x)" | |
| 2690 | by (simp add: path_component_eq_connected_component) | |
| 2691 | ||
| 2692 | lemma locally_path_connected_path_component: | |
| 2693 | "locally path_connected S \<Longrightarrow> locally path_connected (path_component_set S x)" | |
| 2694 | using locally_path_connected_connected_component path_component_eq_connected_component by fastforce | |
| 2695 | ||
| 2696 | lemma open_path_connected_component: | |
| 2697 | fixes S :: "'a :: real_normed_vector set" | |
| 2698 | shows "open S \<Longrightarrow> path_component S x = connected_component S x" | |
| 2699 | by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) | |
| 2700 | ||
| 2701 | lemma open_path_connected_component_set: | |
| 2702 | fixes S :: "'a :: real_normed_vector set" | |
| 2703 | shows "open S \<Longrightarrow> path_component_set S x = connected_component_set S x" | |
| 2704 | by (simp add: open_path_connected_component) | |
| 2705 | ||
| 2706 | proposition locally_connected_quotient_image: | |
| 2707 | assumes lcS: "locally connected S" | |
| 2708 | and oo: "\<And>T. T \<subseteq> f ` S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2709 | \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2710 | openin (top_of_set (f ` S)) T" | 
| 69620 | 2711 | shows "locally connected (f ` S)" | 
| 2712 | proof (clarsimp simp: locally_connected_open_component) | |
| 2713 | fix U C | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2714 | assume opefSU: "openin (top_of_set (f ` S)) U" and "C \<in> components U" | 
| 69620 | 2715 | then have "C \<subseteq> U" "U \<subseteq> f ` S" | 
| 2716 | by (meson in_components_subset openin_imp_subset)+ | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2717 | then have "openin (top_of_set (f ` S)) C \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2718 | openin (top_of_set S) (S \<inter> f -` C)" | 
| 69620 | 2719 | by (auto simp: oo) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2720 | moreover have "openin (top_of_set S) (S \<inter> f -` C)" | 
| 69620 | 2721 | proof (subst openin_subopen, clarify) | 
| 2722 | fix x | |
| 2723 | assume "x \<in> S" "f x \<in> C" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2724 | show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)" | 
| 69620 | 2725 | proof (intro conjI exI) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2726 | show "openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)" | 
| 69620 | 2727 | proof (rule ccontr) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2728 | assume **: "\<not> openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)" | 
| 69620 | 2729 | then have "x \<notin> (S \<inter> f -` U)" | 
| 2730 | using \<open>U \<subseteq> f ` S\<close> opefSU lcS locally_connected_2 oo by blast | |
| 2731 | with ** show False | |
| 2732 | by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) | |
| 2733 | qed | |
| 2734 | next | |
| 2735 | show "x \<in> connected_component_set (S \<inter> f -` U) x" | |
| 2736 | using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by auto | |
| 2737 | next | |
| 2738 | have contf: "continuous_on S f" | |
| 2739 | by (simp add: continuous_on_open oo openin_imp_subset) | |
| 2740 | then have "continuous_on (connected_component_set (S \<inter> f -` U) x) f" | |
| 71769 | 2741 | by (meson connected_component_subset continuous_on_subset inf.boundedE) | 
| 69620 | 2742 | then have "connected (f ` connected_component_set (S \<inter> f -` U) x)" | 
| 2743 | by (rule connected_continuous_image [OF _ connected_connected_component]) | |
| 2744 | moreover have "f ` connected_component_set (S \<inter> f -` U) x \<subseteq> U" | |
| 2745 | using connected_component_in by blast | |
| 2746 |       moreover have "C \<inter> f ` connected_component_set (S \<inter> f -` U) x \<noteq> {}"
 | |
| 2747 | using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by fastforce | |
| 2748 | ultimately have fC: "f ` (connected_component_set (S \<inter> f -` U) x) \<subseteq> C" | |
| 2749 | by (rule components_maximal [OF \<open>C \<in> components U\<close>]) | |
| 2750 | have cUC: "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" | |
| 2751 | using connected_component_subset fC by blast | |
| 2752 | have "connected_component_set (S \<inter> f -` U) x \<subseteq> connected_component_set (S \<inter> f -` C) x" | |
| 2753 | proof - | |
| 2754 |         { assume "x \<in> connected_component_set (S \<inter> f -` U) x"
 | |
| 2755 | then have ?thesis | |
| 2756 | using cUC connected_component_idemp connected_component_mono by blast } | |
| 2757 | then show ?thesis | |
| 2758 | using connected_component_eq_empty by auto | |
| 2759 | qed | |
| 2760 | also have "\<dots> \<subseteq> (S \<inter> f -` C)" | |
| 2761 | by (rule connected_component_subset) | |
| 2762 | finally show "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" . | |
| 2763 | qed | |
| 2764 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2765 | ultimately show "openin (top_of_set (f ` S)) C" | 
| 69620 | 2766 | by metis | 
| 2767 | qed | |
| 2768 | ||
| 2769 | text\<open>The proof resembles that above but is not identical!\<close> | |
| 2770 | proposition locally_path_connected_quotient_image: | |
| 2771 | assumes lcS: "locally path_connected S" | |
| 2772 | and oo: "\<And>T. T \<subseteq> f ` S | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2773 | \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> openin (top_of_set (f ` S)) T" | 
| 69620 | 2774 | shows "locally path_connected (f ` S)" | 
| 2775 | proof (clarsimp simp: locally_path_connected_open_path_component) | |
| 2776 | fix U y | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2777 | assume opefSU: "openin (top_of_set (f ` S)) U" and "y \<in> U" | 
| 69620 | 2778 | then have "path_component_set U y \<subseteq> U" "U \<subseteq> f ` S" | 
| 2779 | by (meson path_component_subset openin_imp_subset)+ | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2780 | then have "openin (top_of_set (f ` S)) (path_component_set U y) \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2781 | openin (top_of_set S) (S \<inter> f -` path_component_set U y)" | 
| 69620 | 2782 | proof - | 
| 2783 | have "path_component_set U y \<subseteq> f ` S" | |
| 2784 | using \<open>U \<subseteq> f ` S\<close> \<open>path_component_set U y \<subseteq> U\<close> by blast | |
| 2785 | then show ?thesis | |
| 2786 | using oo by blast | |
| 2787 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2788 | moreover have "openin (top_of_set S) (S \<inter> f -` path_component_set U y)" | 
| 69620 | 2789 | proof (subst openin_subopen, clarify) | 
| 2790 | fix x | |
| 2791 | assume "x \<in> S" and Uyfx: "path_component U y (f x)" | |
| 2792 | then have "f x \<in> U" | |
| 2793 | using path_component_mem by blast | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2794 | show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)" | 
| 69620 | 2795 | proof (intro conjI exI) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2796 | show "openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)" | 
| 69620 | 2797 | proof (rule ccontr) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2798 | assume **: "\<not> openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)" | 
| 69620 | 2799 | then have "x \<notin> (S \<inter> f -` U)" | 
| 2800 | by (metis (no_types, lifting) \<open>U \<subseteq> f ` S\<close> opefSU lcS oo locally_path_connected_open_path_component) | |
| 2801 | then show False | |
| 2802 | using ** \<open>path_component_set U y \<subseteq> U\<close> \<open>x \<in> S\<close> \<open>path_component U y (f x)\<close> by blast | |
| 2803 | qed | |
| 2804 | next | |
| 2805 | show "x \<in> path_component_set (S \<inter> f -` U) x" | |
| 2806 | by (simp add: \<open>f x \<in> U\<close> \<open>x \<in> S\<close> path_component_refl) | |
| 2807 | next | |
| 2808 | have contf: "continuous_on S f" | |
| 2809 | by (simp add: continuous_on_open oo openin_imp_subset) | |
| 2810 | then have "continuous_on (path_component_set (S \<inter> f -` U) x) f" | |
| 71769 | 2811 | by (meson Int_lower1 continuous_on_subset path_component_subset) | 
| 69620 | 2812 | then have "path_connected (f ` path_component_set (S \<inter> f -` U) x)" | 
| 2813 | by (simp add: path_connected_continuous_image) | |
| 2814 | moreover have "f ` path_component_set (S \<inter> f -` U) x \<subseteq> U" | |
| 2815 | using path_component_mem by fastforce | |
| 2816 | moreover have "f x \<in> f ` path_component_set (S \<inter> f -` U) x" | |
| 2817 | by (force simp: \<open>x \<in> S\<close> \<open>f x \<in> U\<close> path_component_refl_eq) | |
| 2818 | ultimately have "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U (f x)" | |
| 2819 | by (meson path_component_maximal) | |
| 2820 | also have "\<dots> \<subseteq> path_component_set U y" | |
| 2821 | by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) | |
| 2822 | finally have fC: "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U y" . | |
| 2823 | have cUC: "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" | |
| 2824 | using path_component_subset fC by blast | |
| 2825 | have "path_component_set (S \<inter> f -` U) x \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) x" | |
| 2826 | proof - | |
| 2827 | have "\<And>a. path_component_set (path_component_set (S \<inter> f -` U) x) a \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) a" | |
| 2828 | using cUC path_component_mono by blast | |
| 2829 | then show ?thesis | |
| 2830 | using path_component_path_component by blast | |
| 2831 | qed | |
| 2832 | also have "\<dots> \<subseteq> (S \<inter> f -` path_component_set U y)" | |
| 2833 | by (rule path_component_subset) | |
| 2834 | finally show "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" . | |
| 2835 | qed | |
| 2836 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2837 | ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" | 
| 69620 | 2838 | by metis | 
| 2839 | qed | |
| 2840 | ||
| 70136 | 2841 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Components, continuity, openin, closedin\<close> | 
| 69620 | 2842 | |
| 2843 | lemma continuous_on_components_gen: | |
| 2844 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | |
| 71769 | 2845 | assumes "\<And>C. C \<in> components S \<Longrightarrow> | 
| 2846 | openin (top_of_set S) C \<and> continuous_on C f" | |
| 69620 | 2847 | shows "continuous_on S f" | 
| 2848 | proof (clarsimp simp: continuous_openin_preimage_eq) | |
| 2849 | fix t :: "'b set" | |
| 2850 | assume "open t" | |
| 2851 | have *: "S \<inter> f -` t = (\<Union>c \<in> components S. c \<inter> f -` t)" | |
| 2852 | by auto | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2853 | show "openin (top_of_set S) (S \<inter> f -` t)" | 
| 69620 | 2854 | unfolding * using \<open>open t\<close> assms continuous_openin_preimage_gen openin_trans openin_Union by blast | 
| 2855 | qed | |
| 2856 | ||
| 2857 | lemma continuous_on_components: | |
| 2858 | fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" | |
| 71769 | 2859 | assumes "locally connected S " "\<And>C. C \<in> components S \<Longrightarrow> continuous_on C f" | 
| 2860 | shows "continuous_on S f" | |
| 2861 | proof (rule continuous_on_components_gen) | |
| 2862 | fix C | |
| 2863 | assume "C \<in> components S" | |
| 2864 | then show "openin (top_of_set S) C \<and> continuous_on C f" | |
| 2865 | by (simp add: assms openin_components_locally_connected) | |
| 2866 | qed | |
| 69620 | 2867 | |
| 2868 | lemma continuous_on_components_eq: | |
| 2869 | "locally connected S | |
| 2870 | \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" | |
| 2871 | by (meson continuous_on_components continuous_on_subset in_components_subset) | |
| 2872 | ||
| 2873 | lemma continuous_on_components_open: | |
| 2874 | fixes S :: "'a::real_normed_vector set" | |
| 2875 | assumes "open S " | |
| 2876 | "\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f" | |
| 2877 | shows "continuous_on S f" | |
| 2878 | using continuous_on_components open_imp_locally_connected assms by blast | |
| 2879 | ||
| 2880 | lemma continuous_on_components_open_eq: | |
| 2881 | fixes S :: "'a::real_normed_vector set" | |
| 2882 | shows "open S \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" | |
| 2883 | using continuous_on_subset in_components_subset | |
| 2884 | by (blast intro: continuous_on_components_open) | |
| 2885 | ||
| 2886 | lemma closedin_union_complement_components: | |
| 71769 | 2887 | assumes U: "locally connected U" | 
| 2888 | and S: "closedin (top_of_set U) S" | |
| 2889 | and cuS: "c \<subseteq> components(U - S)" | |
| 2890 | shows "closedin (top_of_set U) (S \<union> \<Union>c)" | |
| 69620 | 2891 | proof - | 
| 72372 | 2892 | have di: "(\<And>S T. S \<in> c \<and> T \<in> c' \<Longrightarrow> disjnt S T) \<Longrightarrow> disjnt (\<Union> c) (\<Union> c')" for c' | 
| 69620 | 2893 | by (simp add: disjnt_def) blast | 
| 71769 | 2894 | have "S \<subseteq> U" | 
| 69620 | 2895 | using S closedin_imp_subset by blast | 
| 71769 | 2896 | moreover have "U - S = \<Union>c \<union> \<Union>(components (U - S) - c)" | 
| 69620 | 2897 | by (metis Diff_partition Union_components Union_Un_distrib assms(3)) | 
| 71769 | 2898 | moreover have "disjnt (\<Union>c) (\<Union>(components (U - S) - c))" | 
| 69620 | 2899 | apply (rule di) | 
| 72372 | 2900 | by (metis di DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) | 
| 71769 | 2901 | ultimately have eq: "S \<union> \<Union>c = U - (\<Union>(components(U - S) - c))" | 
| 69620 | 2902 | by (auto simp: disjnt_def) | 
| 71769 | 2903 | have *: "openin (top_of_set U) (\<Union>(components (U - S) - c))" | 
| 2904 | proof (rule openin_Union [OF openin_trans [of "U - S"]]) | |
| 2905 | show "openin (top_of_set (U - S)) T" if "T \<in> components (U - S) - c" for T | |
| 2906 | using that by (simp add: U S locally_diff_closed openin_components_locally_connected) | |
| 2907 | show "openin (top_of_set U) (U - S)" if "T \<in> components (U - S) - c" for T | |
| 2908 | using that by (simp add: openin_diff S) | |
| 2909 | qed | |
| 2910 | have "closedin (top_of_set U) (U - \<Union> (components (U - S) - c))" | |
| 2911 | by (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) | |
| 2912 | then have "openin (top_of_set U) (U - (U - \<Union>(components (U - S) - c)))" | |
| 2913 | by (simp add: openin_diff) | |
| 69620 | 2914 | then show ?thesis | 
| 2915 | by (force simp: eq closedin_def) | |
| 2916 | qed | |
| 2917 | ||
| 2918 | lemma closed_union_complement_components: | |
| 2919 | fixes S :: "'a::real_normed_vector set" | |
| 2920 | assumes S: "closed S" and c: "c \<subseteq> components(- S)" | |
| 2921 | shows "closed(S \<union> \<Union> c)" | |
| 2922 | proof - | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2923 | have "closedin (top_of_set UNIV) (S \<union> \<Union>c)" | 
| 71769 | 2924 | by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_union_complement_components locally_connected_UNIV subtopology_UNIV) | 
| 69620 | 2925 | then show ?thesis by simp | 
| 2926 | qed | |
| 2927 | ||
| 2928 | lemma closedin_Un_complement_component: | |
| 2929 | fixes S :: "'a::real_normed_vector set" | |
| 2930 | assumes u: "locally connected u" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2931 | and S: "closedin (top_of_set u) S" | 
| 69620 | 2932 | and c: " c \<in> components(u - S)" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2933 | shows "closedin (top_of_set u) (S \<union> c)" | 
| 69620 | 2934 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2935 |   have "closedin (top_of_set u) (S \<union> \<Union>{c})"
 | 
| 69620 | 2936 | using c by (blast intro: closedin_union_complement_components [OF u S]) | 
| 2937 | then show ?thesis | |
| 2938 | by simp | |
| 2939 | qed | |
| 2940 | ||
| 2941 | lemma closed_Un_complement_component: | |
| 2942 | fixes S :: "'a::real_normed_vector set" | |
| 2943 | assumes S: "closed S" and c: " c \<in> components(-S)" | |
| 2944 | shows "closed (S \<union> c)" | |
| 2945 | by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component | |
| 2946 | locally_connected_UNIV subtopology_UNIV) | |
| 2947 | ||
| 2948 | ||
| 2949 | subsection\<open>Existence of isometry between subspaces of same dimension\<close> | |
| 2950 | ||
| 2951 | lemma isometry_subset_subspace: | |
| 2952 | fixes S :: "'a::euclidean_space set" | |
| 2953 | and T :: "'b::euclidean_space set" | |
| 2954 | assumes S: "subspace S" | |
| 2955 | and T: "subspace T" | |
| 2956 | and d: "dim S \<le> dim T" | |
| 78457 | 2957 | obtains f where "linear f" "f \<in> S \<rightarrow> T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" | 
| 69620 | 2958 | proof - | 
| 2959 | obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" | |
| 2960 | and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" | |
| 2961 | and "independent B" "finite B" "card B = dim S" "span B = S" | |
| 78516 
56a408fa2440
substantial tidy-up, shortening many proofs
 paulson <lp15@cam.ac.uk> parents: 
78474diff
changeset | 2962 | by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) | 
| 69620 | 2963 | obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" | 
| 2964 | and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" | |
| 2965 | and "independent C" "finite C" "card C = dim T" "span C = T" | |
| 78516 
56a408fa2440
substantial tidy-up, shortening many proofs
 paulson <lp15@cam.ac.uk> parents: 
78474diff
changeset | 2966 | by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) | 
| 69620 | 2967 | obtain fb where "fb ` B \<subseteq> C" "inj_on fb B" | 
| 2968 | by (metis \<open>card B = dim S\<close> \<open>card C = dim T\<close> \<open>finite B\<close> \<open>finite C\<close> card_le_inj d) | |
| 2969 | then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" | |
| 72372 | 2970 | using Corth unfolding pairwise_def inj_on_def | 
| 2971 | by (blast intro: orthogonal_clauses) | |
| 69620 | 2972 | obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" | 
| 2973 | using linear_independent_extend \<open>independent B\<close> by fastforce | |
| 2974 | have "span (f ` B) \<subseteq> span C" | |
| 2975 | by (metis \<open>fb ` B \<subseteq> C\<close> ffb image_cong span_mono) | |
| 2976 | then have "f ` S \<subseteq> T" | |
| 2977 | unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] . | |
| 2978 | have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" | |
| 2979 | using B1 C1 \<open>fb ` B \<subseteq> C\<close> by auto | |
| 2980 | have "norm (f x) = norm x" if "x \<in> S" for x | |
| 2981 | proof - | |
| 2982 | interpret linear f by fact | |
| 2983 | obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" | |
| 2984 | using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce | |
| 2985 | have "norm (f x)^2 = norm (\<Sum>v\<in>B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) | |
| 2986 | also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" | |
| 71769 | 2987 | proof (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) | 
| 2988 | show "pairwise (\<lambda>v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" | |
| 2989 | by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) | |
| 2990 | qed | |
| 69620 | 2991 | also have "\<dots> = norm x ^2" | 
| 2992 | by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) | |
| 2993 | finally show ?thesis | |
| 2994 | by (simp add: norm_eq_sqrt_inner) | |
| 2995 | qed | |
| 2996 | then show ?thesis | |
| 78457 | 2997 | by (meson \<open>f ` S \<subseteq> T\<close> \<open>linear f\<close> image_subset_iff_funcset that) | 
| 69620 | 2998 | qed | 
| 2999 | ||
| 3000 | proposition isometries_subspaces: | |
| 3001 | fixes S :: "'a::euclidean_space set" | |
| 3002 | and T :: "'b::euclidean_space set" | |
| 3003 | assumes S: "subspace S" | |
| 3004 | and T: "subspace T" | |
| 3005 | and d: "dim S = dim T" | |
| 3006 | obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" | |
| 3007 | "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" | |
| 3008 | "\<And>x. x \<in> T \<Longrightarrow> norm(g x) = norm x" | |
| 3009 | "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" | |
| 3010 | "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" | |
| 3011 | proof - | |
| 3012 | obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" | |
| 3013 | and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" | |
| 3014 | and "independent B" "finite B" "card B = dim S" "span B = S" | |
| 78516 
56a408fa2440
substantial tidy-up, shortening many proofs
 paulson <lp15@cam.ac.uk> parents: 
78474diff
changeset | 3015 | by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) | 
| 69620 | 3016 | obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" | 
| 3017 | and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" | |
| 3018 | and "independent C" "finite C" "card C = dim T" "span C = T" | |
| 78516 
56a408fa2440
substantial tidy-up, shortening many proofs
 paulson <lp15@cam.ac.uk> parents: 
78474diff
changeset | 3019 | by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) | 
| 69620 | 3020 | obtain fb where "bij_betw fb B C" | 
| 3021 | by (metis \<open>finite B\<close> \<open>finite C\<close> bij_betw_iff_card \<open>card B = dim S\<close> \<open>card C = dim T\<close> d) | |
| 3022 | then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" | |
| 72372 | 3023 | using Corth unfolding pairwise_def inj_on_def bij_betw_def | 
| 3024 | by (blast intro: orthogonal_clauses) | |
| 69620 | 3025 | obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" | 
| 3026 | using linear_independent_extend \<open>independent B\<close> by fastforce | |
| 3027 | interpret f: linear f by fact | |
| 3028 | define gb where "gb \<equiv> inv_into B fb" | |
| 3029 | then have pairwise_orth_gb: "pairwise (\<lambda>v j. orthogonal (gb v) (gb j)) C" | |
| 72372 | 3030 | using Borth \<open>bij_betw fb B C\<close> unfolding pairwise_def bij_betw_def by force | 
| 69620 | 3031 | obtain g where "linear g" and ggb: "\<And>x. x \<in> C \<Longrightarrow> g x = gb x" | 
| 3032 | using linear_independent_extend \<open>independent C\<close> by fastforce | |
| 3033 | interpret g: linear g by fact | |
| 3034 | have "span (f ` B) \<subseteq> span C" | |
| 3035 | by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on eq_iff ffb image_cong) | |
| 3036 | then have "f ` S \<subseteq> T" | |
| 72372 | 3037 | unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] . | 
| 69620 | 3038 | have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" | 
| 3039 | using B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on by fastforce | |
| 3040 | have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \<in> S" for x | |
| 3041 | proof - | |
| 3042 | obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" | |
| 3043 | using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce | |
| 3044 | have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))" | |
| 3045 | using linear_sum [OF \<open>linear f\<close>] x by auto | |
| 3046 | also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R f v)" | |
| 3047 | by (simp add: f.sum f.scale) | |
| 3048 | also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R fb v)" | |
| 3049 | by (simp add: ffb cong: sum.cong) | |
| 3050 | finally have *: "f x = (\<Sum>v\<in>B. a v *\<^sub>R fb v)" . | |
| 3051 | then have "(norm (f x))\<^sup>2 = (norm (\<Sum>v\<in>B. a v *\<^sub>R fb v))\<^sup>2" by simp | |
| 3052 | also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" | |
| 71769 | 3053 | proof (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) | 
| 3054 | show "pairwise (\<lambda>v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" | |
| 3055 | by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) | |
| 3056 | qed | |
| 69620 | 3057 | also have "\<dots> = (norm x)\<^sup>2" | 
| 3058 | by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) | |
| 3059 | finally show "norm (f x) = norm x" | |
| 3060 | by (simp add: norm_eq_sqrt_inner) | |
| 3061 | have "g (f x) = g (\<Sum>v\<in>B. a v *\<^sub>R fb v)" by (simp add: *) | |
| 3062 | also have "\<dots> = (\<Sum>v\<in>B. g (a v *\<^sub>R fb v))" | |
| 3063 | by (simp add: g.sum g.scale) | |
| 3064 | also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R g (fb v))" | |
| 3065 | by (simp add: g.scale) | |
| 3066 | also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R v)" | |
| 71769 | 3067 | proof (rule sum.cong [OF refl]) | 
| 3068 | show "a x *\<^sub>R g (fb x) = a x *\<^sub>R x" if "x \<in> B" for x | |
| 3069 | using that \<open>bij_betw fb B C\<close> bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce | |
| 3070 | qed | |
| 69620 | 3071 | also have "\<dots> = x" | 
| 3072 | using x by blast | |
| 3073 | finally show "g (f x) = x" . | |
| 3074 | qed | |
| 3075 | have [simp]: "\<And>x. x \<in> C \<Longrightarrow> norm (gb x) = norm x" | |
| 3076 | by (metis B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on gb_def inv_into_into) | |
| 3077 | have g [simp]: "f (g x) = x" if "x \<in> T" for x | |
| 3078 | proof - | |
| 3079 | obtain a where x: "x = (\<Sum>v \<in> C. a v *\<^sub>R v)" | |
| 3080 | using \<open>finite C\<close> \<open>span C = T\<close> \<open>x \<in> T\<close> span_finite by fastforce | |
| 3081 | have "g x = (\<Sum>v \<in> C. g (a v *\<^sub>R v))" | |
| 3082 | by (simp add: x g.sum) | |
| 3083 | also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R g v)" | |
| 3084 | by (simp add: g.scale) | |
| 3085 | also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R gb v)" | |
| 3086 | by (simp add: ggb cong: sum.cong) | |
| 3087 | finally have "f (g x) = f (\<Sum>v\<in>C. a v *\<^sub>R gb v)" by simp | |
| 3088 | also have "\<dots> = (\<Sum>v\<in>C. f (a v *\<^sub>R gb v))" | |
| 3089 | by (simp add: f.scale f.sum) | |
| 3090 | also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R f (gb v))" | |
| 3091 | by (simp add: f.scale f.sum) | |
| 3092 | also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R v)" | |
| 3093 | using \<open>bij_betw fb B C\<close> | |
| 3094 | by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) | |
| 3095 | also have "\<dots> = x" | |
| 3096 | using x by blast | |
| 3097 | finally show "f (g x) = x" . | |
| 3098 | qed | |
| 3099 | have gim: "g ` T = S" | |
| 3100 | by (metis (full_types) S T \<open>f ` S \<subseteq> T\<close> d dim_eq_span dim_image_le f(2) g.linear_axioms | |
| 3101 | image_iff linear_subspace_image span_eq_iff subset_iff) | |
| 3102 | have fim: "f ` S = T" | |
| 3103 | using \<open>g ` T = S\<close> image_iff by fastforce | |
| 3104 | have [simp]: "norm (g x) = norm x" if "x \<in> T" for x | |
| 3105 | using fim that by auto | |
| 3106 | show ?thesis | |
| 71769 | 3107 | by (rule that [OF \<open>linear f\<close> \<open>linear g\<close>]) (simp_all add: fim gim) | 
| 69620 | 3108 | qed | 
| 3109 | ||
| 3110 | corollary isometry_subspaces: | |
| 3111 | fixes S :: "'a::euclidean_space set" | |
| 3112 | and T :: "'b::euclidean_space set" | |
| 3113 | assumes S: "subspace S" | |
| 3114 | and T: "subspace T" | |
| 3115 | and d: "dim S = dim T" | |
| 3116 | obtains f where "linear f" "f ` S = T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" | |
| 3117 | using isometries_subspaces [OF assms] | |
| 3118 | by metis | |
| 3119 | ||
| 3120 | corollary isomorphisms_UNIV_UNIV: | |
| 3121 |   assumes "DIM('M) = DIM('N)"
 | |
| 3122 | obtains f::"'M::euclidean_space \<Rightarrow>'N::euclidean_space" and g | |
| 3123 | where "linear f" "linear g" | |
| 3124 | "\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" | |
| 3125 | "\<And>x. g (f x) = x" "\<And>y. f(g y) = y" | |
| 3126 | using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) | |
| 3127 | ||
| 3128 | lemma homeomorphic_subspaces: | |
| 3129 | fixes S :: "'a::euclidean_space set" | |
| 3130 | and T :: "'b::euclidean_space set" | |
| 3131 | assumes S: "subspace S" | |
| 3132 | and T: "subspace T" | |
| 3133 | and d: "dim S = dim T" | |
| 3134 | shows "S homeomorphic T" | |
| 3135 | proof - | |
| 3136 | obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" | |
| 3137 | "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" | |
| 3138 | by (blast intro: isometries_subspaces [OF assms]) | |
| 3139 | then show ?thesis | |
| 72372 | 3140 | unfolding homeomorphic_def homeomorphism_def | 
| 3141 | apply (rule_tac x=f in exI, rule_tac x=g in exI) | |
| 69620 | 3142 | apply (auto simp: linear_continuous_on linear_conv_bounded_linear) | 
| 3143 | done | |
| 3144 | qed | |
| 3145 | ||
| 3146 | lemma homeomorphic_affine_sets: | |
| 3147 | assumes "affine S" "affine T" "aff_dim S = aff_dim T" | |
| 3148 | shows "S homeomorphic T" | |
| 3149 | proof (cases "S = {} \<or> T = {}")
 | |
| 3150 | case True with assms aff_dim_empty homeomorphic_empty show ?thesis | |
| 3151 | by metis | |
| 3152 | next | |
| 3153 | case False | |
| 3154 | then obtain a b where ab: "a \<in> S" "b \<in> T" by auto | |
| 3155 | then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" | |
| 3156 | using affine_diffs_subspace assms by blast+ | |
| 3157 | have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" | |
| 3158 | using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) | |
| 3159 | have "S homeomorphic ((+) (- a) ` S)" | |
| 69768 | 3160 | by (fact homeomorphic_translation) | 
| 69620 | 3161 | also have "\<dots> homeomorphic ((+) (- b) ` T)" | 
| 3162 | by (rule homeomorphic_subspaces [OF ss dd]) | |
| 3163 | also have "\<dots> homeomorphic T" | |
| 69768 | 3164 | using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) | 
| 69620 | 3165 | finally show ?thesis . | 
| 3166 | qed | |
| 3167 | ||
| 3168 | ||
| 3169 | subsection\<open>Retracts, in a general sense, preserve (co)homotopic triviality)\<close> | |
| 3170 | ||
| 70136 | 3171 | locale\<^marker>\<open>tag important\<close> Retracts = | 
| 77684 | 3172 | fixes S h t k | 
| 3173 | assumes conth: "continuous_on S h" | |
| 3174 | and imh: "h ` S = t" | |
| 69620 | 3175 | and contk: "continuous_on t k" | 
| 78457 | 3176 | and imk: "k \<in> t \<rightarrow> S" | 
| 69620 | 3177 | and idhk: "\<And>y. y \<in> t \<Longrightarrow> h(k y) = y" | 
| 3178 | ||
| 3179 | begin | |
| 3180 | ||
| 3181 | lemma homotopically_trivial_retraction_gen: | |
| 78457 | 3182 | assumes P: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" | 
| 3183 | and Q: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" | |
| 71769 | 3184 | and Qeq: "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | 
| 78457 | 3185 | and hom: "\<And>f g. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f; | 
| 3186 | continuous_on U g; g \<in> U \<rightarrow> S; P g\<rbrakk> | |
| 77684 | 3187 | \<Longrightarrow> homotopic_with_canon P U S f g" | 
| 78457 | 3188 | and contf: "continuous_on U f" and imf: "f \<in> U \<rightarrow> t" and Qf: "Q f" | 
| 3189 | and contg: "continuous_on U g" and img: "g \<in> U \<rightarrow> t" and Qg: "Q g" | |
| 71769 | 3190 | shows "homotopic_with_canon Q U t f g" | 
| 69620 | 3191 | proof - | 
| 71769 | 3192 | have "continuous_on U (k \<circ> f)" | 
| 78457 | 3193 | by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) | 
| 77684 | 3194 | moreover have "(k \<circ> f) ` U \<subseteq> S" | 
| 69620 | 3195 | using imf imk by fastforce | 
| 3196 | moreover have "P (k \<circ> f)" | |
| 3197 | by (simp add: P Qf contf imf) | |
| 71769 | 3198 | moreover have "continuous_on U (k \<circ> g)" | 
| 78457 | 3199 | by (meson contg continuous_on_compose continuous_on_subset contk funcset_image img) | 
| 77684 | 3200 | moreover have "(k \<circ> g) ` U \<subseteq> S" | 
| 69620 | 3201 | using img imk by fastforce | 
| 3202 | moreover have "P (k \<circ> g)" | |
| 3203 | by (simp add: P Qg contg img) | |
| 77684 | 3204 | ultimately have "homotopic_with_canon P U S (k \<circ> f) (k \<circ> g)" | 
| 78457 | 3205 | by (simp add: hom image_subset_iff) | 
| 71769 | 3206 | then have "homotopic_with_canon Q U t (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" | 
| 69620 | 3207 | apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3208 | using Q conth imh by force+ | 
| 69620 | 3209 | then show ?thesis | 
| 72372 | 3210 | proof (rule homotopic_with_eq; simp) | 
| 3211 | show "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | |
| 3212 | using Qeq topspace_euclidean_subtopology by blast | |
| 3213 | show "\<And>x. x \<in> U \<Longrightarrow> f x = h (k (f x))" "\<And>x. x \<in> U \<Longrightarrow> g x = h (k (g x))" | |
| 78457 | 3214 | using idhk imf img by fastforce+ | 
| 72372 | 3215 | qed | 
| 69620 | 3216 | qed | 
| 3217 | ||
| 3218 | lemma homotopically_trivial_retraction_null_gen: | |
| 78457 | 3219 | assumes P: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" | 
| 3220 | and Q: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" | |
| 71769 | 3221 | and Qeq: "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | 
| 78457 | 3222 | and hom: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> | 
| 77684 | 3223 | \<Longrightarrow> \<exists>c. homotopic_with_canon P U S f (\<lambda>x. c)" | 
| 78457 | 3224 | and contf: "continuous_on U f" and imf:"f \<in> U \<rightarrow> t" and Qf: "Q f" | 
| 71769 | 3225 | obtains c where "homotopic_with_canon Q U t f (\<lambda>x. c)" | 
| 69620 | 3226 | proof - | 
| 71769 | 3227 | have feq: "\<And>x. x \<in> U \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto | 
| 3228 | have "continuous_on U (k \<circ> f)" | |
| 78457 | 3229 | by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) | 
| 3230 | moreover have "(k \<circ> f) \<in> U \<rightarrow> S" | |
| 69620 | 3231 | using imf imk by fastforce | 
| 3232 | moreover have "P (k \<circ> f)" | |
| 3233 | by (simp add: P Qf contf imf) | |
| 77684 | 3234 | ultimately obtain c where "homotopic_with_canon P U S (k \<circ> f) (\<lambda>x. c)" | 
| 69620 | 3235 | by (metis hom) | 
| 71769 | 3236 | then have "homotopic_with_canon Q U t (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" | 
| 69620 | 3237 | apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3238 | using Q conth imh by force+ | 
| 71769 | 3239 | then have "homotopic_with_canon Q U t f (\<lambda>x. h c)" | 
| 3240 | proof (rule homotopic_with_eq) | |
| 3241 | show "\<And>x. x \<in> topspace (top_of_set U) \<Longrightarrow> f x = (h \<circ> (k \<circ> f)) x" | |
| 3242 | using feq by auto | |
| 3243 | show "\<And>h k. (\<And>x. x \<in> topspace (top_of_set U) \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | |
| 3244 | using Qeq topspace_euclidean_subtopology by blast | |
| 3245 | qed auto | |
| 69620 | 3246 | then show ?thesis | 
| 71769 | 3247 | using that by blast | 
| 69620 | 3248 | qed | 
| 3249 | ||
| 3250 | lemma cohomotopically_trivial_retraction_gen: | |
| 78457 | 3251 | assumes P: "\<And>f. \<lbrakk>continuous_on t f; f \<in> t \<rightarrow> U; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" | 
| 3252 | and Q: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" | |
| 69620 | 3253 | and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | 
| 78457 | 3254 | and hom: "\<And>f g. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f; | 
| 3255 | continuous_on S g; g \<in> S \<rightarrow> U; P g\<rbrakk> | |
| 77684 | 3256 | \<Longrightarrow> homotopic_with_canon P S U f g" | 
| 78457 | 3257 | and contf: "continuous_on t f" and imf: "f \<in> t \<rightarrow> U" and Qf: "Q f" | 
| 3258 | and contg: "continuous_on t g" and img: "g \<in> t \<rightarrow> U" and Qg: "Q g" | |
| 71769 | 3259 | shows "homotopic_with_canon Q t U f g" | 
| 69620 | 3260 | proof - | 
| 3261 | have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto | |
| 3262 | have geq: "\<And>x. x \<in> t \<Longrightarrow> (g \<circ> h \<circ> k) x = g x" using idhk img by auto | |
| 77684 | 3263 | have "continuous_on S (f \<circ> h)" | 
| 69620 | 3264 | using contf conth continuous_on_compose imh by blast | 
| 78457 | 3265 | moreover have "(f \<circ> h) \<in> S \<rightarrow> U" | 
| 69620 | 3266 | using imf imh by fastforce | 
| 3267 | moreover have "P (f \<circ> h)" | |
| 3268 | by (simp add: P Qf contf imf) | |
| 77684 | 3269 | moreover have "continuous_on S (g \<circ> h)" | 
| 69620 | 3270 | using contg continuous_on_compose continuous_on_subset conth imh by blast | 
| 78457 | 3271 | moreover have "(g \<circ> h) \<in> S \<rightarrow> U" | 
| 69620 | 3272 | using img imh by fastforce | 
| 3273 | moreover have "P (g \<circ> h)" | |
| 3274 | by (simp add: P Qg contg img) | |
| 77684 | 3275 | ultimately have "homotopic_with_canon P S U (f \<circ> h) (g \<circ> h)" | 
| 78457 | 3276 | by (simp add: hom) | 
| 71769 | 3277 | then have "homotopic_with_canon Q t U (f \<circ> h \<circ> k) (g \<circ> h \<circ> k)" | 
| 69620 | 3278 | apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3279 | using Q contk imk by force+ | 
| 69620 | 3280 | then show ?thesis | 
| 71746 | 3281 | proof (rule homotopic_with_eq) | 
| 3282 | show "f x = (f \<circ> h \<circ> k) x" "g x = (g \<circ> h \<circ> k) x" | |
| 3283 | if "x \<in> topspace (top_of_set t)" for x | |
| 3284 | using feq geq that by force+ | |
| 3285 | qed (use Qeq topspace_euclidean_subtopology in blast) | |
| 69620 | 3286 | qed | 
| 3287 | ||
| 3288 | lemma cohomotopically_trivial_retraction_null_gen: | |
| 78457 | 3289 | assumes P: "\<And>f. \<lbrakk>continuous_on t f; f \<in> t \<rightarrow> U; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" | 
| 3290 | and Q: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" | |
| 69620 | 3291 | and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" | 
| 78457 | 3292 | and hom: "\<And>f g. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> | 
| 77684 | 3293 | \<Longrightarrow> \<exists>c. homotopic_with_canon P S U f (\<lambda>x. c)" | 
| 78457 | 3294 | and contf: "continuous_on t f" and imf: "f \<in> t \<rightarrow> U" and Qf: "Q f" | 
| 71769 | 3295 | obtains c where "homotopic_with_canon Q t U f (\<lambda>x. c)" | 
| 69620 | 3296 | proof - | 
| 3297 | have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto | |
| 77684 | 3298 | have "continuous_on S (f \<circ> h)" | 
| 69620 | 3299 | using contf conth continuous_on_compose imh by blast | 
| 78457 | 3300 | moreover have "(f \<circ> h) \<in> S \<rightarrow> U" | 
| 69620 | 3301 | using imf imh by fastforce | 
| 3302 | moreover have "P (f \<circ> h)" | |
| 3303 | by (simp add: P Qf contf imf) | |
| 77684 | 3304 | ultimately obtain c where "homotopic_with_canon P S U (f \<circ> h) (\<lambda>x. c)" | 
| 69620 | 3305 | by (metis hom) | 
| 71769 | 3306 | then have \<section>: "homotopic_with_canon Q t U (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)" | 
| 71746 | 3307 | proof (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) | 
| 77684 | 3308 | show "\<And>h. \<lbrakk>continuous_map (top_of_set S) (top_of_set U) h; P h\<rbrakk> \<Longrightarrow> Q (h \<circ> k)" | 
| 71746 | 3309 | using Q by auto | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3310 | qed (use contk imk in force)+ | 
| 71769 | 3311 | moreover have "homotopic_with_canon Q t U f (\<lambda>x. c)" | 
| 71746 | 3312 | using homotopic_with_eq [OF \<section>] feq Qeq by fastforce | 
| 3313 | ultimately show ?thesis | |
| 3314 | using that by blast | |
| 69620 | 3315 | qed | 
| 3316 | ||
| 3317 | end | |
| 3318 | ||
| 3319 | lemma simply_connected_retraction_gen: | |
| 3320 | shows "\<lbrakk>simply_connected S; continuous_on S h; h ` S = T; | |
| 78457 | 3321 | continuous_on T k; k \<in> T \<rightarrow> S; \<And>y. y \<in> T \<Longrightarrow> h(k y) = y\<rbrakk> | 
| 69620 | 3322 | \<Longrightarrow> simply_connected T" | 
| 3323 | apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) | |
| 3324 | apply (rule Retracts.homotopically_trivial_retraction_gen | |
| 3325 | [of S h _ k _ "\<lambda>p. pathfinish p = pathstart p" "\<lambda>p. pathfinish p = pathstart p"]) | |
| 78457 | 3326 | apply (simp_all add: Retracts_def pathfinish_def pathstart_def image_subset_iff_funcset) | 
| 69620 | 3327 | done | 
| 3328 | ||
| 3329 | lemma homeomorphic_simply_connected: | |
| 3330 | "\<lbrakk>S homeomorphic T; simply_connected S\<rbrakk> \<Longrightarrow> simply_connected T" | |
| 3331 | by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) | |
| 3332 | ||
| 3333 | lemma homeomorphic_simply_connected_eq: | |
| 3334 | "S homeomorphic T \<Longrightarrow> (simply_connected S \<longleftrightarrow> simply_connected T)" | |
| 3335 | by (metis homeomorphic_simply_connected homeomorphic_sym) | |
| 3336 | ||
| 3337 | ||
| 3338 | subsection\<open>Homotopy equivalence\<close> | |
| 3339 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3340 | subsection\<open>Homotopy equivalence of topological spaces.\<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3341 | |
| 70136 | 3342 | definition\<^marker>\<open>tag important\<close> homotopy_equivalent_space | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80052diff
changeset | 3343 | (infix \<open>homotopy'_equivalent'_space\<close> 50) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3344 | where "X homotopy_equivalent_space Y \<equiv> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3345 | (\<exists>f g. continuous_map X Y f \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3346 | continuous_map Y X g \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3347 | homotopic_with (\<lambda>x. True) X X (g \<circ> f) id \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3348 | homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3349 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3350 | lemma homeomorphic_imp_homotopy_equivalent_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3351 | "X homeomorphic_space Y \<Longrightarrow> X homotopy_equivalent_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3352 | unfolding homeomorphic_space_def homotopy_equivalent_space_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3353 | apply (erule ex_forward)+ | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3354 | by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3355 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3356 | lemma homotopy_equivalent_space_refl: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3357 | "X homotopy_equivalent_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3358 | by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3359 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3360 | lemma homotopy_equivalent_space_sym: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3361 | "X homotopy_equivalent_space Y \<longleftrightarrow> Y homotopy_equivalent_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3362 | by (meson homotopy_equivalent_space_def) | 
| 69620 | 3363 | |
| 3364 | lemma homotopy_eqv_trans [trans]: | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3365 | assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3366 | shows "X homotopy_equivalent_space U" | 
| 69620 | 3367 | proof - | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3368 | obtain f1 g1 where f1: "continuous_map X Y f1" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3369 | and g1: "continuous_map Y X g1" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3370 | and hom1: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> f1) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3371 | "homotopic_with (\<lambda>x. True) Y Y (f1 \<circ> g1) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3372 | using 1 by (auto simp: homotopy_equivalent_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3373 | obtain f2 g2 where f2: "continuous_map Y U f2" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3374 | and g2: "continuous_map U Y g2" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3375 | and hom2: "homotopic_with (\<lambda>x. True) Y Y (g2 \<circ> f2) id" | 
| 69620 | 3376 | "homotopic_with (\<lambda>x. True) U U (f2 \<circ> g2) id" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3377 | using 2 by (auto simp: homotopy_equivalent_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3378 | have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> f2 \<circ> f1) (id \<circ> f1)" | 
| 71745 | 3379 | using f1 hom2(1) homotopic_with_compose_continuous_map_right by metis | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3380 | then have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> (f2 \<circ> f1)) (id \<circ> f1)" | 
| 69620 | 3381 | by (simp add: o_assoc) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3382 | then have "homotopic_with (\<lambda>x. True) X X | 
| 69620 | 3383 | (g1 \<circ> (g2 \<circ> (f2 \<circ> f1))) (g1 \<circ> (id \<circ> f1))" | 
| 71745 | 3384 | by (simp add: g1 homotopic_with_compose_continuous_map_left) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3385 | moreover have "homotopic_with (\<lambda>x. True) X X (g1 \<circ> id \<circ> f1) id" | 
| 69620 | 3386 | using hom1 by simp | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3387 | ultimately have SS: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> g2 \<circ> (f2 \<circ> f1)) id" | 
| 71769 | 3388 | by (metis comp_assoc homotopic_with_trans id_comp) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3389 | have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> g1 \<circ> g2) (id \<circ> g2)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3390 | using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3391 | then have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> (g1 \<circ> g2)) (id \<circ> g2)" | 
| 69620 | 3392 | by (simp add: o_assoc) | 
| 3393 | then have "homotopic_with (\<lambda>x. True) U U | |
| 3394 | (f2 \<circ> (f1 \<circ> (g1 \<circ> g2))) (f2 \<circ> (id \<circ> g2))" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3395 | by (simp add: f2 homotopic_with_compose_continuous_map_left) | 
| 69620 | 3396 | moreover have "homotopic_with (\<lambda>x. True) U U (f2 \<circ> id \<circ> g2) id" | 
| 3397 | using hom2 by simp | |
| 3398 | ultimately have UU: "homotopic_with (\<lambda>x. True) U U (f2 \<circ> f1 \<circ> (g1 \<circ> g2)) id" | |
| 71769 | 3399 | by (simp add: fun.map_comp hom2(2) homotopic_with_trans) | 
| 69620 | 3400 | show ?thesis | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3401 | unfolding homotopy_equivalent_space_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3402 | by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3403 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3404 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3405 | lemma deformation_retraction_imp_homotopy_equivalent_space: | 
| 77684 | 3406 | "\<lbrakk>homotopic_with (\<lambda>x. True) X X (S \<circ> r) id; retraction_maps X Y r S\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3407 | \<Longrightarrow> X homotopy_equivalent_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3408 | unfolding homotopy_equivalent_space_def retraction_maps_def | 
| 71769 | 3409 | using homotopic_with_id2 by fastforce | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3410 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3411 | lemma deformation_retract_imp_homotopy_equivalent_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3412 | "\<lbrakk>homotopic_with (\<lambda>x. True) X X r id; retraction_maps X Y r id\<rbrakk> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3413 | \<Longrightarrow> X homotopy_equivalent_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3414 | using deformation_retraction_imp_homotopy_equivalent_space by force | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3415 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3416 | lemma deformation_retract_of_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3417 | "S \<subseteq> topspace X \<and> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3418 | (\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id) \<longleftrightarrow> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3419 | S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` (topspace X) \<subseteq> S)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3420 | proof (cases "S \<subseteq> topspace X") | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3421 | case True | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3422 | moreover have "(\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3423 | \<longleftrightarrow> (S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` topspace X \<subseteq> S))" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3424 | unfolding retract_of_space_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3425 | proof safe | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3426 | fix f r | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3427 | assume f: "homotopic_with (\<lambda>x. True) X X id f" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3428 | and fS: "f ` topspace X \<subseteq> S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3429 | and r: "continuous_map X (subtopology X S) r" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3430 | and req: "\<forall>x\<in>S. r x = x" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3431 | show "\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3432 | proof (intro exI conjI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3433 | have "homotopic_with (\<lambda>x. True) X X f r" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3434 | proof (rule homotopic_with_eq) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3435 | show "homotopic_with (\<lambda>x. True) X X (r \<circ> f) (r \<circ> id)" | 
| 71745 | 3436 | by (metis continuous_map_into_fulltopology f homotopic_with_compose_continuous_map_left homotopic_with_symD r) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3437 | show "f x = (r \<circ> f) x" if "x \<in> topspace X" for x | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3438 | using that fS req by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3439 | qed auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3440 | then show "homotopic_with (\<lambda>x. True) X X id r" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3441 | by (rule homotopic_with_trans [OF f]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3442 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3443 | show "retraction_maps X (subtopology X S) r id" | 
| 71172 | 3444 | by (simp add: r req retraction_maps_def) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3445 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3446 | qed (use True in \<open>auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\<close>) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3447 | ultimately show ?thesis by simp | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3448 | qed (auto simp: retract_of_space_def retraction_maps_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3449 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3450 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3451 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3452 | subsection\<open>Contractible spaces\<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3453 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3454 | text\<open>The definition (which agrees with "contractible" on subsets of Euclidean space) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3455 | is a little cryptic because we don't in fact assume that the constant "a" is in the space. | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3456 | This forces the convention that the empty space / set is contractible, avoiding some special cases. \<close> | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3457 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3458 | definition contractible_space where | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3459 | "contractible_space X \<equiv> \<exists>a. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3460 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3461 | lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \<longleftrightarrow> contractible S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3462 | by (auto simp: contractible_space_def contractible_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3463 | |
| 78336 | 3464 | lemma contractible_space_empty [simp]: | 
| 3465 | "contractible_space trivial_topology" | |
| 71745 | 3466 | unfolding contractible_space_def homotopic_with_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3467 | apply (rule_tac x=undefined in exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3468 | apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else undefined" in exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3469 | apply (auto simp: continuous_map_on_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3470 | done | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3471 | |
| 78336 | 3472 | lemma contractible_space_singleton [simp]: | 
| 3473 |   "contractible_space (discrete_topology{a})"
 | |
| 71745 | 3474 | unfolding contractible_space_def homotopic_with_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3475 | apply (rule_tac x=a in exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3476 | apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else a" in exI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3477 | apply (auto intro: continuous_map_eq [where f = "\<lambda>z. a"]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3478 | done | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3479 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3480 | lemma contractible_space_subset_singleton: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3481 |    "topspace X \<subseteq> {a} \<Longrightarrow> contractible_space X"
 | 
| 78336 | 3482 | by (metis contractible_space_empty contractible_space_singleton null_topspace_iff_trivial subset_singletonD subtopology_eq_discrete_topology_sing) | 
| 3483 | ||
| 3484 | lemma contractible_space_subtopology_singleton [simp]: | |
| 3485 |    "contractible_space (subtopology X {a})"
 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3486 | by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3487 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3488 | lemma contractible_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3489 | "contractible_space X \<longleftrightarrow> | 
| 78336 | 3490 | X = trivial_topology \<or> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3491 | (\<exists>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))" | 
| 78336 | 3492 | proof (cases "X = trivial_topology") | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3493 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3494 | then show ?thesis | 
| 71745 | 3495 | using homotopic_with_imp_continuous_maps by (fastforce simp: contractible_space_def) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3496 | qed (simp add: contractible_space_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3497 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3498 | lemma contractible_imp_path_connected_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3499 | assumes "contractible_space X" shows "path_connected_space X" | 
| 78336 | 3500 | proof (cases "X = trivial_topology") | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3501 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3502 | have *: "path_connected_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3503 |     if "a \<in> topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3504 | and h: "\<forall>x. h (0, x) = x" "\<forall>x. h (1, x) = a" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3505 | for a and h :: "real \<times> 'a \<Rightarrow> 'a" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3506 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3507 | have "path_component_of X b a" if "b \<in> topspace X" for b | 
| 71745 | 3508 | unfolding path_component_of_def | 
| 3509 | proof (intro exI conjI) | |
| 3510 | let ?g = "h \<circ> (\<lambda>x. (x,b))" | |
| 3511 | show "pathin X ?g" | |
| 3512 | unfolding pathin_def | |
| 71769 | 3513 | proof (rule continuous_map_compose [OF _ conth]) | 
| 3514 |         show "continuous_map (top_of_set {0..1}) (prod_topology (top_of_set {0..1}) X) (\<lambda>x. (x, b))"
 | |
| 3515 | using that by (auto intro!: continuous_intros) | |
| 3516 | qed | |
| 71745 | 3517 | qed (use h in auto) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3518 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3519 | by (metis path_component_of_equiv path_connected_space_iff_path_component) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3520 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3521 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3522 | using assms False by (auto simp: contractible_space homotopic_with_def *) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3523 | qed (simp add: path_connected_space_topspace_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3524 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3525 | lemma contractible_imp_connected_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3526 | "contractible_space X \<Longrightarrow> connected_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3527 | by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3528 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3529 | lemma contractible_space_alt: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3530 | "contractible_space X \<longleftrightarrow> (\<forall>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))" (is "?lhs = ?rhs") | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3531 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3532 | assume X: ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3533 | then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3534 | by (auto simp: contractible_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3535 | show ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3536 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3537 | show "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. b)" if "b \<in> topspace X" for b | 
| 71745 | 3538 | proof (rule homotopic_with_trans [OF a]) | 
| 3539 | show "homotopic_with (\<lambda>x. True) X X (\<lambda>x. a) (\<lambda>x. b)" | |
| 3540 | using homotopic_constant_maps path_connected_space_imp_path_component_of | |
| 78336 | 3541 | by (metis X a contractible_imp_path_connected_space homotopic_with_sym homotopic_with_trans path_component_of_equiv that) | 
| 71745 | 3542 | qed | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3543 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3544 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3545 | assume R: ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3546 | then show ?lhs | 
| 78336 | 3547 | using contractible_space_def by fastforce | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3548 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3549 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3550 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3551 | lemma compose_const [simp]: "f \<circ> (\<lambda>x. a) = (\<lambda>x. f a)" "(\<lambda>x. a) \<circ> g = (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3552 | by (simp_all add: o_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3553 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3554 | lemma nullhomotopic_through_contractible_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3555 | assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3556 | obtains c where "homotopic_with (\<lambda>h. True) X Z (g \<circ> f) (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3557 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3558 | obtain b where b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3559 | using Y by (auto simp: contractible_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3560 | show thesis | 
| 71745 | 3561 | using homotopic_with_compose_continuous_map_right | 
| 3562 | [OF homotopic_with_compose_continuous_map_left [OF b g] f] | |
| 78474 | 3563 | by (force simp: that) | 
| 69620 | 3564 | qed | 
| 3565 | ||
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3566 | lemma nullhomotopic_into_contractible_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3567 | assumes f: "continuous_map X Y f" and Y: "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3568 | obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3569 | using nullhomotopic_through_contractible_space [OF f _ Y] | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3570 | by (metis continuous_map_id id_comp) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3571 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3572 | lemma nullhomotopic_from_contractible_space: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3573 | assumes f: "continuous_map X Y f" and X: "contractible_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3574 | obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3575 | using nullhomotopic_through_contractible_space [OF _ f X] | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3576 | by (metis comp_id continuous_map_id) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3577 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3578 | lemma homotopy_dominated_contractibility: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3579 | assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3580 | and hom: "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id" and X: "contractible_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3581 | shows "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3582 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3583 | obtain c where c: "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3584 | using nullhomotopic_from_contractible_space [OF f X] . | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3585 | have "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) (\<lambda>x. c)" | 
| 71745 | 3586 | using homotopic_with_compose_continuous_map_right [OF c g] by fastforce | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3587 | then have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3588 | using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3589 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3590 | unfolding contractible_space_def .. | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3591 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3592 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3593 | lemma homotopy_equivalent_space_contractibility: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3594 | "X homotopy_equivalent_space Y \<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3595 | unfolding homotopy_equivalent_space_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3596 | by (blast intro: homotopy_dominated_contractibility) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3597 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3598 | lemma homeomorphic_space_contractibility: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3599 | "X homeomorphic_space Y | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3600 | \<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3601 | by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3602 | |
| 78127 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3603 | lemma homotopic_through_contractible_space: | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3604 | "continuous_map X Y f \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3605 | continuous_map X Y f' \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3606 | continuous_map Y Z g \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3607 | continuous_map Y Z g' \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3608 | contractible_space Y \<and> path_connected_space Z | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3609 | \<Longrightarrow> homotopic_with (\<lambda>h. True) X Z (g \<circ> f) (g' \<circ> f')" | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3610 | using nullhomotopic_through_contractible_space [of X Y f Z g] | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3611 | using nullhomotopic_through_contractible_space [of X Y f' Z g'] | 
| 78336 | 3612 | by (smt (verit) continuous_map_const homotopic_constant_maps homotopic_with_imp_continuous_maps | 
| 3613 | homotopic_with_symD homotopic_with_trans path_connected_space_imp_path_component_of) | |
| 78127 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3614 | |
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3615 | lemma homotopic_from_contractible_space: | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3616 | "continuous_map X Y f \<and> continuous_map X Y g \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3617 | contractible_space X \<and> path_connected_space Y | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3618 | \<Longrightarrow> homotopic_with (\<lambda>x. True) X Y f g" | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3619 | by (metis comp_id continuous_map_id homotopic_through_contractible_space) | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3620 | |
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3621 | lemma homotopic_into_contractible_space: | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3622 | "continuous_map X Y f \<and> continuous_map X Y g \<and> | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3623 | contractible_space Y | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3624 | \<Longrightarrow> homotopic_with (\<lambda>x. True) X Y f g" | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3625 | by (metis continuous_map_id contractible_imp_path_connected_space homotopic_through_contractible_space id_comp) | 
| 
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
77929diff
changeset | 3626 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3627 | lemma contractible_eq_homotopy_equivalent_singleton_subtopology: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3628 | "contractible_space X \<longleftrightarrow> | 
| 78336 | 3629 |         X = trivial_topology \<or> (\<exists>a \<in> topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs")
 | 
| 3630 | proof (cases "X = trivial_topology") | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3631 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3632 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3633 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3634 | assume ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3635 | then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3636 | by (auto simp: contractible_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3637 | then have "a \<in> topspace X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3638 | by (metis False continuous_map_const homotopic_with_imp_continuous_maps) | 
| 72372 | 3639 |     then have "homotopic_with (\<lambda>x. True) (subtopology X {a}) (subtopology X {a}) id (\<lambda>x. a)"
 | 
| 3640 | using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3641 |     then have "X homotopy_equivalent_space subtopology X {a}"
 | 
| 72372 | 3642 | unfolding homotopy_equivalent_space_def using \<open>a \<in> topspace X\<close> | 
| 3643 | by (metis (full_types) a comp_id continuous_map_const continuous_map_id_subt empty_subsetI homotopic_with_symD | |
| 3644 | id_comp insertI1 insert_subset topspace_subtopology_subset) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3645 | with \<open>a \<in> topspace X\<close> show ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3646 | by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3647 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3648 | assume ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3649 | then show ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3650 | by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3651 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3652 | qed (simp add: contractible_space_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3653 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3654 | lemma contractible_space_retraction_map_image: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3655 | assumes "retraction_map X Y f" and X: "contractible_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3656 | shows "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3657 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3658 | obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\<forall>y \<in> topspace Y. f(g y) = y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3659 | using assms by (auto simp: retraction_map_def retraction_maps_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3660 | obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3661 | using X by (auto simp: contractible_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3662 | have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. f a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3663 | proof (rule homotopic_with_eq) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3664 | show "homotopic_with (\<lambda>x. True) Y Y (f \<circ> id \<circ> g) (f \<circ> (\<lambda>x. a) \<circ> g)" | 
| 71745 | 3665 | using f g a homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right by metis | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3666 | qed (use fg in auto) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3667 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3668 | unfolding contractible_space_def by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3669 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3670 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3671 | lemma contractible_space_prod_topology: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3672 | "contractible_space(prod_topology X Y) \<longleftrightarrow> | 
| 78336 | 3673 | X = trivial_topology \<or> Y = trivial_topology \<or> contractible_space X \<and> contractible_space Y" | 
| 3674 | proof (cases "X = trivial_topology \<or> Y = trivial_topology") | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3675 | case True | 
| 78336 | 3676 | then have "(prod_topology X Y) = trivial_topology" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3677 | by simp | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3678 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3679 | by (auto simp: contractible_space_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3680 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3681 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3682 | have "contractible_space(prod_topology X Y) \<longleftrightarrow> contractible_space X \<and> contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3683 | proof safe | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3684 | assume XY: "contractible_space (prod_topology X Y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3685 | with False have "retraction_map (prod_topology X Y) X fst" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3686 | by (auto simp: contractible_space False retraction_map_fst) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3687 | then show "contractible_space X" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3688 | by (rule contractible_space_retraction_map_image [OF _ XY]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3689 | have "retraction_map (prod_topology X Y) Y snd" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3690 | using False XY by (auto simp: contractible_space False retraction_map_snd) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3691 | then show "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3692 | by (rule contractible_space_retraction_map_image [OF _ XY]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3693 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3694 | assume "contractible_space X" and "contractible_space Y" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3695 | with False obtain a b | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3696 | where "a \<in> topspace X" and a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3697 | and "b \<in> topspace Y" and b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3698 | by (auto simp: contractible_space) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3699 | with False show "contractible_space (prod_topology X Y)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3700 | apply (simp add: contractible_space) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3701 | apply (rule_tac x=a in bexI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3702 | apply (rule_tac x=b in bexI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3703 | using homotopic_with_prod_topology [OF a b] | 
| 70033 | 3704 | apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) | 
| 3705 | apply auto | |
| 3706 | done | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3707 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3708 | with False show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3709 | by auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3710 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3711 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3712 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3713 | lemma contractible_space_product_topology: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3714 | "contractible_space(product_topology X I) \<longleftrightarrow> | 
| 78336 | 3715 | (product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. contractible_space(X i))" | 
| 3716 | proof (cases "(product_topology X I) = trivial_topology") | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3717 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3718 | have 1: "contractible_space (X i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3719 | if XI: "contractible_space (product_topology X I)" and "i \<in> I" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3720 | for i | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3721 | proof (rule contractible_space_retraction_map_image [OF _ XI]) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3722 | show "retraction_map (product_topology X I) (X i) (\<lambda>x. x i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3723 | using False by (simp add: retraction_map_product_projection \<open>i \<in> I\<close>) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3724 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3725 | have 2: "contractible_space (product_topology X I)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3726 | if "x \<in> topspace (product_topology X I)" and cs: "\<forall>i\<in>I. contractible_space (X i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3727 | for x :: "'a \<Rightarrow> 'b" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3728 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3729 | obtain f where f: "\<And>i. i\<in>I \<Longrightarrow> homotopic_with (\<lambda>x. True) (X i) (X i) id (\<lambda>x. f i)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3730 | using cs unfolding contractible_space_def by metis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3731 | have "homotopic_with (\<lambda>x. True) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3732 | (product_topology X I) (product_topology X I) id (\<lambda>x. restrict f I)" | 
| 71172 | 3733 | by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3734 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3735 | by (auto simp: contractible_space_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3736 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3737 | show ?thesis | 
| 78336 | 3738 | using False 1 2 by (meson equals0I subtopology_eq_discrete_topology_empty) | 
| 3739 | qed auto | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3740 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3741 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3742 | lemma contractible_space_subtopology_euclideanreal [simp]: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3743 | "contractible_space(subtopology euclideanreal S) \<longleftrightarrow> is_interval S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3744 | (is "?lhs = ?rhs") | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3745 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3746 | assume ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3747 | then have "path_connectedin (subtopology euclideanreal S) S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3748 | using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3749 | by (simp add: contractible_imp_path_connected) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3750 | then show ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3751 | by (simp add: is_interval_path_connected_1) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3752 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3753 | assume ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3754 | then have "convex S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3755 | by (simp add: is_interval_convex_1) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3756 | show ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3757 |   proof (cases "S = {}")
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3758 | case False | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3759 | then obtain z where "z \<in> S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3760 | by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3761 | show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3762 | unfolding contractible_space_def homotopic_with_def | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3763 | proof (intro exI conjI allI) | 
| 71745 | 3764 | note \<section> = convexD [OF \<open>convex S\<close>, simplified] | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3765 |       show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S)
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3766 | (\<lambda>(t,x). (1 - t) * x + t * z)" | 
| 71745 | 3767 | using \<open>z \<in> S\<close> | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 3768 | by (auto simp: case_prod_unfold intro!: continuous_intros \<section>) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3769 | qed auto | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3770 | qed (simp add: contractible_space_empty) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3771 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3772 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3773 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3774 | corollary contractible_space_euclideanreal: "contractible_space euclideanreal" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3775 | proof - | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3776 | have "contractible_space (subtopology euclideanreal UNIV)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3777 | using contractible_space_subtopology_euclideanreal by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3778 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3779 | by simp | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3780 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3781 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3782 | |
| 70136 | 3783 | abbreviation\<^marker>\<open>tag important\<close> homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool" | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80052diff
changeset | 3784 | (infix \<open>homotopy'_eqv\<close> 50) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3785 | where "S homotopy_eqv T \<equiv> top_of_set S homotopy_equivalent_space top_of_set T" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3786 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3787 | lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \<Longrightarrow> S homotopy_eqv T" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3788 | unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def | 
| 71769 | 3789 | by (metis continuous_map_subtopology_eu homotopic_with_id2 openin_imp_subset openin_subtopology_self topspace_euclidean_subtopology) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3790 | |
| 69620 | 3791 | lemma homotopy_eqv_inj_linear_image: | 
| 3792 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 3793 | assumes "linear f" "inj f" | |
| 3794 | shows "(f ` S) homotopy_eqv S" | |
| 71745 | 3795 | by (metis assms homeomorphic_sym homeomorphic_imp_homotopy_eqv linear_homeomorphic_image) | 
| 69620 | 3796 | |
| 3797 | lemma homotopy_eqv_translation: | |
| 3798 | fixes S :: "'a::real_normed_vector set" | |
| 3799 | shows "(+) a ` S homotopy_eqv S" | |
| 71745 | 3800 | using homeomorphic_imp_homotopy_eqv homeomorphic_translation homeomorphic_sym by blast | 
| 69620 | 3801 | |
| 3802 | lemma homotopy_eqv_homotopic_triviality_imp: | |
| 3803 | fixes S :: "'a::real_normed_vector set" | |
| 3804 | and T :: "'b::real_normed_vector set" | |
| 3805 | and U :: "'c::real_normed_vector set" | |
| 3806 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3807 | and f: "continuous_on U f" "f \<in> U \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3808 | and g: "continuous_on U g" "g \<in> U \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3809 | and homUS: "\<And>f g. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3810 | continuous_on U g; g \<in> U \<rightarrow> S\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3811 | \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3812 | shows "homotopic_with_canon (\<lambda>x. True) U T f g" | 
| 69620 | 3813 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3814 | obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3815 | and k: "continuous_on T k" "k \<in> T \<rightarrow> S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3816 | and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3817 | "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3818 | using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3819 | have "homotopic_with_canon (\<lambda>f. True) U S (k \<circ> f) (k \<circ> g)" | 
| 71745 | 3820 | proof (rule homUS) | 
| 3821 | show "continuous_on U (k \<circ> f)" "continuous_on U (k \<circ> g)" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3822 | using continuous_on_compose continuous_on_subset f g k by (metis funcset_image)+ | 
| 71745 | 3823 | qed (use f g k in \<open>(force simp: o_def)+\<close> ) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3824 | then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3825 | by (simp add: h homotopic_with_compose_continuous_map_left image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3826 | moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> f) (id \<circ> f)" | 
| 71745 | 3827 | by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom f) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3828 | moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> g) (id \<circ> g)" | 
| 71745 | 3829 | by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom g) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3830 | ultimately show "homotopic_with_canon (\<lambda>x. True) U T f g" | 
| 71745 | 3831 | unfolding o_assoc | 
| 3832 | by (metis homotopic_with_trans homotopic_with_sym id_comp) | |
| 69620 | 3833 | qed | 
| 3834 | ||
| 3835 | lemma homotopy_eqv_homotopic_triviality: | |
| 3836 | fixes S :: "'a::real_normed_vector set" | |
| 3837 | and T :: "'b::real_normed_vector set" | |
| 3838 | and U :: "'c::real_normed_vector set" | |
| 3839 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3840 | shows "(\<forall>f g. continuous_on U f \<and> f \<in> U \<rightarrow> S \<and> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3841 | continuous_on U g \<and> g \<in> U \<rightarrow> S | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3842 | \<longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g) \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3843 | (\<forall>f g. continuous_on U f \<and> f \<in> U \<rightarrow> T \<and> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3844 | continuous_on U g \<and> g \<in> U \<rightarrow> T | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3845 | \<longrightarrow> homotopic_with_canon (\<lambda>x. True) U T f g)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3846 | (is "?lhs = ?rhs") | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3847 | proof | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3848 | assume ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3849 | then show ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3850 | by (metis assms homotopy_eqv_homotopic_triviality_imp) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3851 | next | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3852 | assume ?rhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3853 | moreover | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3854 | have "T homotopy_eqv S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3855 | using assms homotopy_equivalent_space_sym by blast | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3856 | ultimately show ?lhs | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3857 | by (blast intro: homotopy_eqv_homotopic_triviality_imp) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3858 | qed | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3859 | |
| 69620 | 3860 | |
| 3861 | lemma homotopy_eqv_cohomotopic_triviality_null_imp: | |
| 3862 | fixes S :: "'a::real_normed_vector set" | |
| 3863 | and T :: "'b::real_normed_vector set" | |
| 3864 | and U :: "'c::real_normed_vector set" | |
| 3865 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3866 | and f: "continuous_on T f" "f \<in> T \<rightarrow> U" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3867 | and homSU: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3868 | \<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3869 | obtains c where "homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)" | 
| 69620 | 3870 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3871 | obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3872 | and k: "continuous_on T k" "k \<in> T \<rightarrow> S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3873 | and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3874 | "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3875 | using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3876 | obtain c where "homotopic_with_canon (\<lambda>x. True) S U (f \<circ> h) (\<lambda>x. c)" | 
| 71745 | 3877 | proof (rule exE [OF homSU]) | 
| 3878 | show "continuous_on S (f \<circ> h)" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3879 | by (metis continuous_on_compose continuous_on_subset f h funcset_image) | 
| 71745 | 3880 | qed (use f h in force) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3881 | then have "homotopic_with_canon (\<lambda>x. True) T U ((f \<circ> h) \<circ> k) ((\<lambda>x. c) \<circ> k)" | 
| 71745 | 3882 | by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3883 | moreover have "homotopic_with_canon (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))" | 
| 71745 | 3884 | by (rule homotopic_with_compose_continuous_left [where Y=T]) | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 3885 | (use f in \<open>auto simp: hom homotopic_with_symD\<close>) | 
| 69620 | 3886 | ultimately show ?thesis | 
| 78474 | 3887 | using that homotopic_with_trans by (fastforce simp: o_def) | 
| 69620 | 3888 | qed | 
| 3889 | ||
| 3890 | lemma homotopy_eqv_cohomotopic_triviality_null: | |
| 3891 | fixes S :: "'a::real_normed_vector set" | |
| 3892 | and T :: "'b::real_normed_vector set" | |
| 3893 | and U :: "'c::real_normed_vector set" | |
| 3894 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3895 | shows "(\<forall>f. continuous_on S f \<and> f \<in> S \<rightarrow> U | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3896 | \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c))) \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3897 | (\<forall>f. continuous_on T f \<and> f \<in> T \<rightarrow> U | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3898 | \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)))" | 
| 71745 | 3899 | by (rule iffI; metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym) | 
| 3900 | ||
| 3901 | text \<open>Similar to the proof above\<close> | |
| 69620 | 3902 | lemma homotopy_eqv_homotopic_triviality_null_imp: | 
| 3903 | fixes S :: "'a::real_normed_vector set" | |
| 3904 | and T :: "'b::real_normed_vector set" | |
| 3905 | and U :: "'c::real_normed_vector set" | |
| 3906 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3907 | and f: "continuous_on U f" "f \<in> U \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3908 | and homSU: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3909 | \<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3910 | shows "\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)" | 
| 69620 | 3911 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3912 | obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3913 | and k: "continuous_on T k" "k \<in> T \<rightarrow> S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3914 | and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3915 | "homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3916 | using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3917 | obtain c::'a where "homotopic_with_canon (\<lambda>x. True) U S (k \<circ> f) (\<lambda>x. c)" | 
| 71745 | 3918 | proof (rule exE [OF homSU [of "k \<circ> f"]]) | 
| 3919 | show "continuous_on U (k \<circ> f)" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3920 | using continuous_on_compose continuous_on_subset f k by (metis funcset_image) | 
| 71745 | 3921 | qed (use f k in force) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3922 | then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" | 
| 71745 | 3923 | by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3924 | moreover have "homotopic_with_canon (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)" | 
| 71745 | 3925 | by (rule homotopic_with_compose_continuous_right [where X=T]) | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 3926 | (use f in \<open>auto simp: hom homotopic_with_symD\<close>) | 
| 69620 | 3927 | ultimately show ?thesis | 
| 78474 | 3928 | using homotopic_with_trans by (fastforce simp: o_def) | 
| 69620 | 3929 | qed | 
| 3930 | ||
| 3931 | lemma homotopy_eqv_homotopic_triviality_null: | |
| 3932 | fixes S :: "'a::real_normed_vector set" | |
| 3933 | and T :: "'b::real_normed_vector set" | |
| 3934 | and U :: "'c::real_normed_vector set" | |
| 3935 | assumes "S homotopy_eqv T" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3936 | shows "(\<forall>f. continuous_on U f \<and> f \<in> U \<rightarrow> S | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3937 | \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c))) \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78127diff
changeset | 3938 | (\<forall>f. continuous_on U f \<and> f \<in> U \<rightarrow> T | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3939 | \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)))" | 
| 71745 | 3940 | by (rule iffI; metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym) | 
| 69620 | 3941 | |
| 3942 | lemma homotopy_eqv_contractible_sets: | |
| 3943 | fixes S :: "'a::real_normed_vector set" | |
| 3944 | and T :: "'b::real_normed_vector set" | |
| 3945 |   assumes "contractible S" "contractible T" "S = {} \<longleftrightarrow> T = {}"
 | |
| 3946 | shows "S homotopy_eqv T" | |
| 3947 | proof (cases "S = {}")
 | |
| 3948 | case True with assms show ?thesis | |
| 78336 | 3949 | using homeomorphic_imp_homotopy_eqv by fastforce | 
| 69620 | 3950 | next | 
| 3951 | case False | |
| 3952 | with assms obtain a b where "a \<in> S" "b \<in> T" | |
| 3953 | by auto | |
| 3954 | then show ?thesis | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3955 | unfolding homotopy_equivalent_space_def | 
| 72372 | 3956 | apply (rule_tac x="\<lambda>x. b" in exI, rule_tac x="\<lambda>x. a" in exI) | 
| 3957 | apply (intro assms conjI continuous_on_id' homotopic_into_contractible; force) | |
| 69620 | 3958 | done | 
| 3959 | qed | |
| 3960 | ||
| 3961 | lemma homotopy_eqv_empty1 [simp]: | |
| 3962 | fixes S :: "'a::real_normed_vector set" | |
| 71769 | 3963 |   shows "S homotopy_eqv ({}::'b::real_normed_vector set) \<longleftrightarrow> S = {}" (is "?lhs = ?rhs")
 | 
| 3964 | proof | |
| 3965 | assume ?lhs then show ?rhs | |
| 3966 | by (metis continuous_map_subtopology_eu empty_iff equalityI homotopy_equivalent_space_def image_subset_iff subsetI) | |
| 78336 | 3967 | qed (use homeomorphic_imp_homotopy_eqv in force) | 
| 69620 | 3968 | |
| 3969 | lemma homotopy_eqv_empty2 [simp]: | |
| 3970 | fixes S :: "'a::real_normed_vector set" | |
| 3971 |   shows "({}::'b::real_normed_vector set) homotopy_eqv S \<longleftrightarrow> S = {}"
 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3972 | using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast | 
| 69620 | 3973 | |
| 3974 | lemma homotopy_eqv_contractibility: | |
| 3975 | fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" | |
| 3976 | shows "S homotopy_eqv T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 3977 | by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility) | 
| 69620 | 3978 | |
| 3979 | lemma homotopy_eqv_sing: | |
| 3980 | fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" | |
| 3981 |   shows "S homotopy_eqv {a} \<longleftrightarrow> S \<noteq> {} \<and> contractible S"
 | |
| 78336 | 3982 | by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets homotopy_eqv_empty2) | 
| 69620 | 3983 | |
| 3984 | lemma homeomorphic_contractible_eq: | |
| 3985 | fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" | |
| 3986 | shows "S homeomorphic T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" | |
| 3987 | by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) | |
| 3988 | ||
| 3989 | lemma homeomorphic_contractible: | |
| 3990 | fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" | |
| 3991 | shows "\<lbrakk>contractible S; S homeomorphic T\<rbrakk> \<Longrightarrow> contractible T" | |
| 3992 | by (metis homeomorphic_contractible_eq) | |
| 3993 | ||
| 3994 | ||
| 70136 | 3995 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Misc other results\<close> | 
| 69620 | 3996 | |
| 3997 | lemma bounded_connected_Compl_real: | |
| 3998 | fixes S :: "real set" | |
| 3999 | assumes "bounded S" and conn: "connected(- S)" | |
| 4000 |     shows "S = {}"
 | |
| 4001 | proof - | |
| 4002 | obtain a b where "S \<subseteq> box a b" | |
| 4003 | by (meson assms bounded_subset_box_symmetric) | |
| 4004 | then have "a \<notin> S" "b \<notin> S" | |
| 4005 | by auto | |
| 4006 | then have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> - S" | |
| 4007 | by (meson Compl_iff conn connected_iff_interval) | |
| 4008 | then show ?thesis | |
| 4009 | using \<open>S \<subseteq> box a b\<close> by auto | |
| 4010 | qed | |
| 4011 | ||
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69768diff
changeset | 4012 | corollary bounded_path_connected_Compl_real: | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69768diff
changeset | 4013 | fixes S :: "real set" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69768diff
changeset | 4014 |   assumes "bounded S" "path_connected(- S)" shows "S = {}"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69768diff
changeset | 4015 | by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69768diff
changeset | 4016 | |
| 69620 | 4017 | lemma bounded_connected_Compl_1: | 
| 4018 |   fixes S :: "'a::{euclidean_space} set"
 | |
| 4019 |   assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1"
 | |
| 4020 |     shows "S = {}"
 | |
| 4021 | proof - | |
| 4022 |   have "DIM('a) = DIM(real)"
 | |
| 4023 | by (simp add: "1") | |
| 4024 | then obtain f::"'a \<Rightarrow> real" and g | |
| 71769 | 4025 | where "linear f" "\<And>x. norm(f x) = norm x" and fg: "\<And>x. g(f x) = x" "\<And>y. f(g y) = y" | 
| 69620 | 4026 | by (rule isomorphisms_UNIV_UNIV) blast | 
| 4027 | with \<open>bounded S\<close> have "bounded (f ` S)" | |
| 4028 | using bounded_linear_image linear_linear by blast | |
| 71769 | 4029 | have "bij f" by (metis fg bijI') | 
| 69620 | 4030 | have "connected (f ` (-S))" | 
| 4031 | using connected_linear_image assms \<open>linear f\<close> by blast | |
| 4032 | moreover have "f ` (-S) = - (f ` S)" | |
| 71769 | 4033 | by (simp add: \<open>bij f\<close> bij_image_Compl_eq) | 
| 69620 | 4034 | finally have "connected (- (f ` S))" | 
| 4035 | by simp | |
| 4036 |   then have "f ` S = {}"
 | |
| 4037 | using \<open>bounded (f ` S)\<close> bounded_connected_Compl_real by blast | |
| 4038 | then show ?thesis | |
| 4039 | by blast | |
| 4040 | qed | |
| 4041 | ||
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 4042 | lemma connected_card_eq_iff_nontrivial: | 
| 69620 | 4043 | fixes S :: "'a::metric_space set" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 4044 |   shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> \<not>(\<exists>a. S \<subseteq> {a})"
 | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 4045 | by (metis connected_uncountable finite.emptyI finite.insertI rev_finite_subset singleton_iff subsetI uncountable_infinite) | 
| 69620 | 4046 | |
| 4047 | lemma connected_finite_iff_sing: | |
| 4048 | fixes S :: "'a::metric_space set" | |
| 4049 | assumes "connected S" | |
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 4050 |   shows "finite S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})" 
 | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
76874diff
changeset | 4051 | using assms connected_uncountable countable_finite by blast | 
| 69620 | 4052 | |
| 70136 | 4053 | subsection\<^marker>\<open>tag unimportant\<close>\<open> Some simple positive connection theorems\<close> | 
| 69620 | 4054 | |
| 4055 | proposition path_connected_convex_diff_countable: | |
| 4056 | fixes U :: "'a::euclidean_space set" | |
| 4057 | assumes "convex U" "\<not> collinear U" "countable S" | |
| 4058 | shows "path_connected(U - S)" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4059 | proof (clarsimp simp: path_connected_def) | 
| 69620 | 4060 | fix a b | 
| 4061 | assume "a \<in> U" "a \<notin> S" "b \<in> U" "b \<notin> S" | |
| 4062 | let ?m = "midpoint a b" | |
| 4063 | show "\<exists>g. path g \<and> path_image g \<subseteq> U - S \<and> pathstart g = a \<and> pathfinish g = b" | |
| 4064 | proof (cases "a = b") | |
| 4065 | case True | |
| 4066 | then show ?thesis | |
| 4067 | by (metis DiffI \<open>a \<in> U\<close> \<open>a \<notin> S\<close> path_component_def path_component_refl) | |
| 4068 | next | |
| 4069 | case False | |
| 4070 | then have "a \<noteq> ?m" "b \<noteq> ?m" | |
| 4071 | using midpoint_eq_endpoint by fastforce+ | |
| 4072 | have "?m \<in> U" | |
| 4073 | using \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>convex U\<close> convex_contains_segment by force | |
| 4074 |     obtain c where "c \<in> U" and nc_abc: "\<not> collinear {a,b,c}"
 | |
| 4075 | by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>\<not> collinear U\<close> collinear_triples insert_absorb) | |
| 4076 |     have ncoll_mca: "\<not> collinear {?m,c,a}"
 | |
| 4077 | by (metis (full_types) \<open>a \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) | |
| 4078 |     have ncoll_mcb: "\<not> collinear {?m,c,b}"
 | |
| 4079 | by (metis (full_types) \<open>b \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) | |
| 4080 | have "c \<noteq> ?m" | |
| 4081 | by (metis collinear_midpoint insert_commute nc_abc) | |
| 4082 | then have "closed_segment ?m c \<subseteq> U" | |
| 4083 | by (simp add: \<open>c \<in> U\<close> \<open>?m \<in> U\<close> \<open>convex U\<close> closed_segment_subset) | |
| 4084 | then obtain z where z: "z \<in> closed_segment ?m c" | |
| 4085 |                     and disjS: "(closed_segment a z \<union> closed_segment z b) \<inter> S = {}"
 | |
| 4086 | proof - | |
| 4087 |       have False if "closed_segment ?m c \<subseteq> {z. (closed_segment a z \<union> closed_segment z b) \<inter> S \<noteq> {}}"
 | |
| 4088 | proof - | |
| 4089 | have closb: "closed_segment ?m c \<subseteq> | |
| 4090 |                  {z \<in> closed_segment ?m c. closed_segment a z \<inter> S \<noteq> {}} \<union> {z \<in> closed_segment ?m c. closed_segment z b \<inter> S \<noteq> {}}"
 | |
| 4091 | using that by blast | |
| 4092 |         have *: "countable {z \<in> closed_segment ?m c. closed_segment z u \<inter> S \<noteq> {}}"
 | |
| 4093 |           if "u \<in> U" "u \<notin> S" and ncoll: "\<not> collinear {?m, c, u}" for u
 | |
| 4094 | proof - | |
| 4095 | have **: False if x1: "x1 \<in> closed_segment ?m c" and x2: "x2 \<in> closed_segment ?m c" | |
| 4096 | and "x1 \<noteq> x2" "x1 \<noteq> u" | |
| 4097 | and w: "w \<in> closed_segment x1 u" "w \<in> closed_segment x2 u" | |
| 4098 | and "w \<in> S" for x1 x2 w | |
| 4099 | proof - | |
| 4100 |             have "x1 \<in> affine hull {?m,c}" "x2 \<in> affine hull {?m,c}"
 | |
| 4101 | using segment_as_ball x1 x2 by auto | |
| 4102 |             then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}"
 | |
| 4103 | by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) | |
| 4104 |             have "\<not> collinear {x1, u, x2}"
 | |
| 4105 | proof | |
| 4106 |               assume "collinear {x1, u, x2}"
 | |
| 4107 |               then have "collinear {?m, c, u}"
 | |
| 4108 | by (metis (full_types) \<open>c \<noteq> ?m\<close> coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \<open>x1 \<noteq> x2\<close>) | |
| 4109 | with ncoll show False .. | |
| 4110 | qed | |
| 4111 |             then have "closed_segment x1 u \<inter> closed_segment u x2 = {u}"
 | |
| 4112 | by (blast intro!: Int_closed_segment) | |
| 4113 | then have "w = u" | |
| 4114 | using closed_segment_commute w by auto | |
| 4115 | show ?thesis | |
| 4116 | using \<open>u \<notin> S\<close> \<open>w = u\<close> that(7) by auto | |
| 4117 | qed | |
| 4118 |           then have disj: "disjoint ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}))"
 | |
| 4119 | by (fastforce simp: pairwise_def disjnt_def) | |
| 4120 |           have cou: "countable ((\<Union>z \<in> closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})"
 | |
| 4121 | apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) | |
| 4122 | apply (rule countable_subset [OF _ \<open>countable S\<close>], auto) | |
| 4123 | done | |
| 4124 | define f where "f \<equiv> \<lambda>X. (THE z. z \<in> closed_segment ?m c \<and> X = closed_segment z u \<inter> S)" | |
| 4125 | show ?thesis | |
| 4126 | proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) | |
| 4127 | fix x | |
| 4128 |             assume x: "x \<in> closed_segment ?m c" "closed_segment x u \<inter> S \<noteq> {}"
 | |
| 4129 |             show "x \<in> f ` ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})"
 | |
| 4130 | proof (rule_tac x="closed_segment x u \<inter> S" in image_eqI) | |
| 4131 | show "x = f (closed_segment x u \<inter> S)" | |
| 71769 | 4132 | unfolding f_def | 
| 4133 | by (rule the_equality [symmetric]) (use x in \<open>auto dest: **\<close>) | |
| 69620 | 4134 | qed (use x in auto) | 
| 4135 | qed | |
| 4136 | qed | |
| 4137 | have "uncountable (closed_segment ?m c)" | |
| 4138 | by (metis \<open>c \<noteq> ?m\<close> uncountable_closed_segment) | |
| 4139 | then show False | |
| 4140 | using closb * [OF \<open>a \<in> U\<close> \<open>a \<notin> S\<close> ncoll_mca] * [OF \<open>b \<in> U\<close> \<open>b \<notin> S\<close> ncoll_mcb] | |
| 72372 | 4141 | by (simp add: closed_segment_commute countable_subset) | 
| 69620 | 4142 | qed | 
| 4143 | then show ?thesis | |
| 4144 | by (force intro: that) | |
| 4145 | qed | |
| 4146 | show ?thesis | |
| 4147 | proof (intro exI conjI) | |
| 4148 | have "path_image (linepath a z +++ linepath z b) \<subseteq> U" | |
| 4149 | by (metis \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>closed_segment ?m c \<subseteq> U\<close> z \<open>convex U\<close> closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) | |
| 4150 | with disjS show "path_image (linepath a z +++ linepath z b) \<subseteq> U - S" | |
| 4151 | by (force simp: path_image_join) | |
| 4152 | qed auto | |
| 4153 | qed | |
| 4154 | qed | |
| 4155 | ||
| 4156 | ||
| 4157 | corollary connected_convex_diff_countable: | |
| 4158 | fixes U :: "'a::euclidean_space set" | |
| 4159 | assumes "convex U" "\<not> collinear U" "countable S" | |
| 4160 | shows "connected(U - S)" | |
| 4161 | by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) | |
| 4162 | ||
| 4163 | lemma path_connected_punctured_convex: | |
| 4164 | assumes "convex S" and aff: "aff_dim S \<noteq> 1" | |
| 4165 |     shows "path_connected(S - {a})"
 | |
| 4166 | proof - | |
| 4167 | consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \<ge> 2" | |
| 4168 | using assms aff_dim_geq [of S] by linarith | |
| 4169 | then show ?thesis | |
| 4170 | proof cases | |
| 4171 | assume "aff_dim S = -1" | |
| 4172 | then show ?thesis | |
| 4173 | by (metis aff_dim_empty empty_Diff path_connected_empty) | |
| 4174 | next | |
| 4175 | assume "aff_dim S = 0" | |
| 4176 | then show ?thesis | |
| 4177 | by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) | |
| 4178 | next | |
| 4179 | assume ge2: "aff_dim S \<ge> 2" | |
| 4180 | then have "\<not> collinear S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4181 | proof (clarsimp simp: collinear_affine_hull) | 
| 69620 | 4182 | fix u v | 
| 4183 |       assume "S \<subseteq> affine hull {u, v}"
 | |
| 4184 |       then have "aff_dim S \<le> aff_dim {u, v}"
 | |
| 4185 | by (metis (no_types) aff_dim_affine_hull aff_dim_subset) | |
| 4186 | with ge2 show False | |
| 4187 | by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) | |
| 4188 | qed | |
| 71769 | 4189 |     moreover have "countable {a}"
 | 
| 69620 | 4190 | by simp | 
| 71769 | 4191 | ultimately show ?thesis | 
| 4192 | by (metis path_connected_convex_diff_countable [OF \<open>convex S\<close>]) | |
| 69620 | 4193 | qed | 
| 4194 | qed | |
| 4195 | ||
| 4196 | lemma connected_punctured_convex: | |
| 4197 |   shows "\<lbrakk>convex S; aff_dim S \<noteq> 1\<rbrakk> \<Longrightarrow> connected(S - {a})"
 | |
| 4198 | using path_connected_imp_connected path_connected_punctured_convex by blast | |
| 4199 | ||
| 4200 | lemma path_connected_complement_countable: | |
| 4201 | fixes S :: "'a::euclidean_space set" | |
| 4202 |   assumes "2 \<le> DIM('a)" "countable S"
 | |
| 4203 | shows "path_connected(- S)" | |
| 4204 | proof - | |
| 71769 | 4205 | have "\<not> collinear (UNIV::'a set)" | 
| 69620 | 4206 | using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) | 
| 71769 | 4207 | then have "path_connected(UNIV - S)" | 
| 4208 | by (simp add: \<open>countable S\<close> path_connected_convex_diff_countable) | |
| 69620 | 4209 | then show ?thesis | 
| 4210 | by (simp add: Compl_eq_Diff_UNIV) | |
| 4211 | qed | |
| 4212 | ||
| 4213 | proposition path_connected_openin_diff_countable: | |
| 4214 | fixes S :: "'a::euclidean_space set" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4215 | assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4216 | and "\<not> collinear S" "countable T" | 
| 4217 | shows "path_connected(S - T)" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4218 | proof (clarsimp simp: path_connected_component) | 
| 69620 | 4219 | fix x y | 
| 4220 | assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T" | |
| 4221 | show "path_component (S - T) x y" | |
| 4222 | proof (rule connected_equivalence_relation_gen [OF \<open>connected S\<close>, where P = "\<lambda>x. x \<notin> T"]) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4223 | show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (top_of_set S) U" and "x \<in> U" for U x | 
| 69620 | 4224 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4225 | have "openin (top_of_set (affine hull S)) U" | 
| 69620 | 4226 | using opeU ope openin_trans by blast | 
| 4227 | with \<open>x \<in> U\<close> obtain r where Usub: "U \<subseteq> affine hull S" and "r > 0" | |
| 4228 | and subU: "ball x r \<inter> affine hull S \<subseteq> U" | |
| 4229 | by (auto simp: openin_contains_ball) | |
| 4230 | with \<open>x \<in> U\<close> have x: "x \<in> ball x r \<inter> affine hull S" | |
| 4231 | by auto | |
| 4232 |       have "\<not> S \<subseteq> {x}"
 | |
| 4233 | using \<open>\<not> collinear S\<close> collinear_subset by blast | |
| 4234 | then obtain x' where "x' \<noteq> x" "x' \<in> S" | |
| 4235 | by blast | |
| 4236 | obtain y where y: "y \<noteq> x" "y \<in> ball x r \<inter> affine hull S" | |
| 4237 | proof | |
| 4238 | show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \<noteq> x" | |
| 4239 | using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> by auto | |
| 4240 | show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \<in> ball x r \<inter> affine hull S" | |
| 4241 | using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> \<open>x' \<in> S\<close> x | |
| 4242 | by (simp add: dist_norm mem_affine_3_minus hull_inc) | |
| 4243 | qed | |
| 4244 | have "convex (ball x r \<inter> affine hull S)" | |
| 4245 | by (simp add: affine_imp_convex convex_Int) | |
| 4246 | with x y subU have "uncountable U" | |
| 4247 | by (meson countable_subset uncountable_convex) | |
| 4248 | then have "\<not> U \<subseteq> T" | |
| 4249 | using \<open>countable T\<close> countable_subset by blast | |
| 4250 | then show ?thesis by blast | |
| 4251 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4252 | show "\<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and> | 
| 69620 | 4253 | (\<forall>x\<in>U. \<forall>y\<in>U. x \<notin> T \<and> y \<notin> T \<longrightarrow> path_component (S - T) x y)" | 
| 4254 | if "x \<in> S" for x | |
| 4255 | proof - | |
| 4256 | obtain r where Ssub: "S \<subseteq> affine hull S" and "r > 0" | |
| 4257 | and subS: "ball x r \<inter> affine hull S \<subseteq> S" | |
| 4258 | using ope \<open>x \<in> S\<close> by (auto simp: openin_contains_ball) | |
| 4259 | then have conv: "convex (ball x r \<inter> affine hull S)" | |
| 4260 | by (simp add: affine_imp_convex convex_Int) | |
| 4261 | have "\<not> aff_dim (affine hull S) \<le> 1" | |
| 4262 | using \<open>\<not> collinear S\<close> collinear_aff_dim by auto | |
| 72372 | 4263 | then have "\<not> aff_dim (ball x r \<inter> affine hull S) \<le> 1" | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72372diff
changeset | 4264 | by (metis (no_types, opaque_lifting) aff_dim_convex_Int_open IntI open_ball \<open>0 < r\<close> aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) | 
| 69620 | 4265 | then have "\<not> collinear (ball x r \<inter> affine hull S)" | 
| 72372 | 4266 | by (simp add: collinear_aff_dim) | 
| 69620 | 4267 | then have *: "path_connected ((ball x r \<inter> affine hull S) - T)" | 
| 4268 | by (rule path_connected_convex_diff_countable [OF conv _ \<open>countable T\<close>]) | |
| 4269 | have ST: "ball x r \<inter> affine hull S - T \<subseteq> S - T" | |
| 4270 | using subS by auto | |
| 4271 | show ?thesis | |
| 4272 | proof (intro exI conjI) | |
| 4273 | show "x \<in> ball x r \<inter> affine hull S" | |
| 4274 | using \<open>x \<in> S\<close> \<open>r > 0\<close> by (simp add: hull_inc) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4275 | have "openin (top_of_set (affine hull S)) (ball x r \<inter> affine hull S)" | 
| 69620 | 4276 | by (subst inf.commute) (simp add: openin_Int_open) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4277 | then show "openin (top_of_set S) (ball x r \<inter> affine hull S)" | 
| 69620 | 4278 | by (rule openin_subset_trans [OF _ subS Ssub]) | 
| 4279 | qed (use * path_component_trans in \<open>auto simp: path_connected_component path_component_of_subset [OF ST]\<close>) | |
| 4280 | qed | |
| 4281 | qed (use xy path_component_trans in auto) | |
| 4282 | qed | |
| 4283 | ||
| 4284 | corollary connected_openin_diff_countable: | |
| 4285 | fixes S :: "'a::euclidean_space set" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4286 | assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4287 | and "\<not> collinear S" "countable T" | 
| 4288 | shows "connected(S - T)" | |
| 4289 | by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) | |
| 4290 | ||
| 4291 | corollary path_connected_open_diff_countable: | |
| 4292 | fixes S :: "'a::euclidean_space set" | |
| 4293 |   assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T"
 | |
| 4294 | shows "path_connected(S - T)" | |
| 4295 | proof (cases "S = {}")
 | |
| 4296 | case True | |
| 4297 | then show ?thesis | |
| 71172 | 4298 | by (simp) | 
| 69620 | 4299 | next | 
| 4300 | case False | |
| 4301 | show ?thesis | |
| 4302 | proof (rule path_connected_openin_diff_countable) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4303 | show "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4304 | by (simp add: assms hull_subset open_subset) | 
| 4305 | show "\<not> collinear S" | |
| 4306 | using assms False by (simp add: collinear_aff_dim aff_dim_open) | |
| 4307 | qed (simp_all add: assms) | |
| 4308 | qed | |
| 4309 | ||
| 4310 | corollary connected_open_diff_countable: | |
| 4311 | fixes S :: "'a::euclidean_space set" | |
| 4312 |   assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T"
 | |
| 4313 | shows "connected(S - T)" | |
| 4314 | by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) | |
| 4315 | ||
| 4316 | ||
| 4317 | ||
| 70136 | 4318 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Self-homeomorphisms shuffling points about\<close> | 
| 4319 | ||
| 4320 | subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close> | |
| 69620 | 4321 | |
| 4322 | lemma homeomorphism_moving_point_1: | |
| 4323 | fixes a :: "'a::euclidean_space" | |
| 4324 | assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" | |
| 4325 | obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" | |
| 4326 | "f a = u" "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = x" | |
| 4327 | proof - | |
| 4328 | have nou: "norm (u - a) < r" and "u \<in> T" | |
| 4329 | using u by (auto simp: dist_norm norm_minus_commute) | |
| 4330 | then have "0 < r" | |
| 4331 | by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) | |
| 4332 | define f where "f \<equiv> \<lambda>x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" | |
| 4333 | have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" | |
| 4334 | and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a | |
| 4335 | proof - | |
| 4336 | have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" | |
| 4337 | using eq by (simp add: algebra_simps) | |
| 4338 | then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" | |
| 4339 | by (metis diff_divide_distrib) | |
| 4340 | also have "\<dots> \<le> norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" | |
| 4341 | using norm_triangle_ineq by blast | |
| 4342 | also have "\<dots> = norm y + (norm x - norm y) * (norm u / r)" | |
| 4343 | using yx \<open>r > 0\<close> | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 4344 | by (simp add: field_split_simps) | 
| 69620 | 4345 | also have "\<dots> < norm y + (norm x - norm y) * 1" | 
| 71769 | 4346 | proof (subst add_less_cancel_left) | 
| 4347 | show "(norm x - norm y) * (norm u / r) < (norm x - norm y) * 1" | |
| 4348 | proof (rule mult_strict_left_mono) | |
| 4349 | show "norm u / r < 1" | |
| 4350 | using \<open>0 < r\<close> divide_less_eq_1_pos nou by blast | |
| 4351 | qed (simp add: yx) | |
| 4352 | qed | |
| 69620 | 4353 | also have "\<dots> = norm x" | 
| 4354 | by simp | |
| 4355 | finally show False by simp | |
| 4356 | qed | |
| 4357 | have "inj f" | |
| 4358 | unfolding f_def | |
| 4359 | proof (clarsimp simp: inj_on_def) | |
| 4360 | fix x y | |
| 4361 | assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = | |
| 4362 | (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" | |
| 4363 | then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" | |
| 4364 | by (auto simp: algebra_simps) | |
| 4365 | show "x=y" | |
| 4366 | proof (cases "norm (x - a) = norm (y - a)") | |
| 4367 | case True | |
| 4368 | then show ?thesis | |
| 4369 | using eq by auto | |
| 4370 | next | |
| 4371 | case False | |
| 4372 | then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" | |
| 4373 | by linarith | |
| 4374 | then have "False" | |
| 4375 | proof cases | |
| 4376 | case 1 show False | |
| 4377 | using * [OF _ nou 1] eq by simp | |
| 4378 | next | |
| 4379 | case 2 with * [OF eq nou] show False | |
| 4380 | by auto | |
| 4381 | qed | |
| 4382 | then show "x=y" .. | |
| 4383 | qed | |
| 4384 | qed | |
| 4385 | then have inj_onf: "inj_on f (cball a r \<inter> T)" | |
| 4386 | using inj_on_Int by fastforce | |
| 4387 | have contf: "continuous_on (cball a r \<inter> T) f" | |
| 4388 | unfolding f_def using \<open>0 < r\<close> by (intro continuous_intros) blast | |
| 4389 | have fim: "f ` (cball a r \<inter> T) = cball a r \<inter> T" | |
| 4390 | proof | |
| 4391 | have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> r" if "norm y \<le> r" "norm u < r" for y u::'a | |
| 4392 | proof - | |
| 4393 | have "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> norm y + norm((1 - norm y / r) *\<^sub>R u)" | |
| 4394 | using norm_triangle_ineq by blast | |
| 4395 | also have "\<dots> = norm y + abs(1 - norm y / r) * norm u" | |
| 4396 | by simp | |
| 4397 | also have "\<dots> \<le> r" | |
| 4398 | proof - | |
| 4399 | have "(r - norm u) * (r - norm y) \<ge> 0" | |
| 4400 | using that by auto | |
| 4401 | then have "r * norm u + r * norm y \<le> r * r + norm u * norm y" | |
| 4402 | by (simp add: algebra_simps) | |
| 4403 | then show ?thesis | |
| 4404 | using that \<open>0 < r\<close> by (simp add: abs_if field_simps) | |
| 4405 | qed | |
| 4406 | finally show ?thesis . | |
| 4407 | qed | |
| 4408 | have "f ` (cball a r) \<subseteq> cball a r" | |
| 72372 | 4409 | using * nou | 
| 4410 | apply (clarsimp simp: dist_norm norm_minus_commute f_def) | |
| 4411 | by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute) | |
| 69620 | 4412 | moreover have "f ` T \<subseteq> T" | 
| 4413 | unfolding f_def using \<open>affine T\<close> \<open>a \<in> T\<close> \<open>u \<in> T\<close> | |
| 4414 | by (force simp: add.commute mem_affine_3_minus) | |
| 4415 | ultimately show "f ` (cball a r \<inter> T) \<subseteq> cball a r \<inter> T" | |
| 4416 | by blast | |
| 4417 | next | |
| 4418 | show "cball a r \<inter> T \<subseteq> f ` (cball a r \<inter> T)" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4419 | proof (clarsimp simp: dist_norm norm_minus_commute) | 
| 69620 | 4420 | fix x | 
| 4421 | assume x: "norm (x - a) \<le> r" and "x \<in> T" | |
| 4422 |       have "\<exists>v \<in> {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \<bullet> 1 = 0"
 | |
| 4423 | by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4424 | then obtain v where "0 \<le> v" "v \<le> 1" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4425 | and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" | 
| 69620 | 4426 | by auto | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4427 | then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4428 | by (simp add: field_simps norm_minus_commute) | 
| 69620 | 4429 | show "x \<in> f ` (cball a r \<inter> T)" | 
| 4430 | proof (rule image_eqI) | |
| 4431 | show "x = f (x - v *\<^sub>R (u - a))" | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4432 | using \<open>r > 0\<close> v by (simp add: f_def) (simp add: field_simps) | 
| 69620 | 4433 | have "x - v *\<^sub>R (u - a) \<in> cball a r" | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4434 | using \<open>r > 0\<close>\<open>0 \<le> v\<close> | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 4435 | by (simp add: dist_norm n) | 
| 69620 | 4436 | moreover have "x - v *\<^sub>R (u - a) \<in> T" | 
| 71172 | 4437 | by (simp add: f_def \<open>u \<in> T\<close> \<open>x \<in> T\<close> assms mem_affine_3_minus2) | 
| 69620 | 4438 | ultimately show "x - v *\<^sub>R (u - a) \<in> cball a r \<inter> T" | 
| 4439 | by blast | |
| 4440 | qed | |
| 4441 | qed | |
| 4442 | qed | |
| 71769 | 4443 | have "compact (cball a r \<inter> T)" | 
| 4444 | by (simp add: affine_closed compact_Int_closed \<open>affine T\<close>) | |
| 4445 | then obtain g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" | |
| 4446 | by (metis homeomorphism_compact [OF _ contf fim inj_onf]) | |
| 4447 | then show thesis | |
| 4448 | apply (rule_tac f=f in that) | |
| 4449 | using \<open>r > 0\<close> by (simp_all add: f_def dist_norm norm_minus_commute) | |
| 69620 | 4450 | qed | 
| 4451 | ||
| 70136 | 4452 | corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_2: | 
| 69620 | 4453 | fixes a :: "'a::euclidean_space" | 
| 4454 | assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" | |
| 4455 | obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" | |
| 4456 | "f u = v" "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" | |
| 4457 | proof - | |
| 4458 | have "0 < r" | |
| 4459 | by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) | |
| 4460 | obtain f1 g1 where hom1: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f1 g1" | |
| 4461 | and "f1 a = u" and f1: "\<And>x. x \<in> sphere a r \<Longrightarrow> f1 x = x" | |
| 4462 | using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u] by blast | |
| 4463 | obtain f2 g2 where hom2: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f2 g2" | |
| 4464 | and "f2 a = v" and f2: "\<And>x. x \<in> sphere a r \<Longrightarrow> f2 x = x" | |
| 4465 | using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> v] by blast | |
| 4466 | show ?thesis | |
| 4467 | proof | |
| 4468 | show "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) (f2 \<circ> g1) (f1 \<circ> g2)" | |
| 4469 | by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) | |
| 4470 | have "g1 u = a" | |
| 4471 | using \<open>0 < r\<close> \<open>f1 a = u\<close> assms hom1 homeomorphism_apply1 by fastforce | |
| 4472 | then show "(f2 \<circ> g1) u = v" | |
| 4473 | by (simp add: \<open>f2 a = v\<close>) | |
| 4474 | show "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> (f2 \<circ> g1) x = x" | |
| 4475 | using f1 f2 hom1 homeomorphism_apply1 by fastforce | |
| 4476 | qed | |
| 4477 | qed | |
| 4478 | ||
| 4479 | ||
| 70136 | 4480 | corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_3: | 
| 69620 | 4481 | fixes a :: "'a::euclidean_space" | 
| 4482 | assumes "affine T" "a \<in> T" and ST: "ball a r \<inter> T \<subseteq> S" "S \<subseteq> T" | |
| 4483 | and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" | |
| 4484 | obtains f g where "homeomorphism S S f g" | |
| 4485 |                     "f u = v" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T"
 | |
| 4486 | proof - | |
| 4487 | obtain f g where hom: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" | |
| 4488 | and "f u = v" and fid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" | |
| 4489 | using homeomorphism_moving_point_2 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u v] by blast | |
| 4490 | have gid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> g x = x" | |
| 4491 | using fid hom homeomorphism_apply1 by fastforce | |
| 4492 | define ff where "ff \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then f x else x" | |
| 4493 | define gg where "gg \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then g x else x" | |
| 4494 | show ?thesis | |
| 4495 | proof | |
| 4496 | show "homeomorphism S S ff gg" | |
| 4497 | proof (rule homeomorphismI) | |
| 4498 | have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) ff" | |
| 71769 | 4499 | unfolding ff_def | 
| 72372 | 4500 | using homeomorphism_cont1 [OF hom] | 
| 4501 | by (intro continuous_on_cases) (auto simp: affine_closed \<open>affine T\<close> fid) | |
| 69620 | 4502 | then show "continuous_on S ff" | 
| 71769 | 4503 | by (rule continuous_on_subset) (use ST in auto) | 
| 69620 | 4504 | have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) gg" | 
| 71769 | 4505 | unfolding gg_def | 
| 72372 | 4506 | using homeomorphism_cont2 [OF hom] | 
| 4507 | by (intro continuous_on_cases) (auto simp: affine_closed \<open>affine T\<close> gid) | |
| 69620 | 4508 | then show "continuous_on S gg" | 
| 71769 | 4509 | by (rule continuous_on_subset) (use ST in auto) | 
| 69620 | 4510 | show "ff ` S \<subseteq> S" | 
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4511 | proof (clarsimp simp: ff_def) | 
| 69620 | 4512 | fix x | 
| 4513 | assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" | |
| 4514 | then have "f x \<in> cball a r \<inter> T" | |
| 4515 | using homeomorphism_image1 [OF hom] by force | |
| 4516 | then show "f x \<in> S" | |
| 4517 | using ST(1) \<open>x \<in> T\<close> gid hom homeomorphism_def x by fastforce | |
| 4518 | qed | |
| 4519 | show "gg ` S \<subseteq> S" | |
| 77929 
48aa9928f090
Numerous significant simplifications
 paulson <lp15@cam.ac.uk> parents: 
77684diff
changeset | 4520 | proof (clarsimp simp: gg_def) | 
| 69620 | 4521 | fix x | 
| 4522 | assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" | |
| 4523 | then have "g x \<in> cball a r \<inter> T" | |
| 4524 | using homeomorphism_image2 [OF hom] by force | |
| 4525 | then have "g x \<in> ball a r" | |
| 4526 | using homeomorphism_apply2 [OF hom] | |
| 4527 | by (metis Diff_Diff_Int Diff_iff \<open>x \<in> T\<close> cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) | |
| 4528 | then show "g x \<in> S" | |
| 4529 | using ST(1) \<open>g x \<in> cball a r \<inter> T\<close> by force | |
| 4530 | qed | |
| 4531 | show "\<And>x. x \<in> S \<Longrightarrow> gg (ff x) = x" | |
| 71769 | 4532 | unfolding ff_def gg_def | 
| 69620 | 4533 | using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] | 
| 72372 | 4534 | by simp (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) | 
| 69620 | 4535 | show "\<And>x. x \<in> S \<Longrightarrow> ff (gg x) = x" | 
| 71769 | 4536 | unfolding ff_def gg_def | 
| 69620 | 4537 | using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] | 
| 72372 | 4538 | by simp (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) | 
| 69620 | 4539 | qed | 
| 4540 | show "ff u = v" | |
| 4541 | using u by (auto simp: ff_def \<open>f u = v\<close>) | |
| 4542 |     show "{x. \<not> (ff x = x \<and> gg x = x)} \<subseteq> ball a r \<inter> T"
 | |
| 4543 | by (auto simp: ff_def gg_def) | |
| 4544 | qed | |
| 4545 | qed | |
| 4546 | ||
| 4547 | ||
| 70136 | 4548 | proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point: | 
| 69620 | 4549 | fixes a :: "'a::euclidean_space" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4550 | assumes ope: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4551 | and "S \<subseteq> T" | 
| 4552 | and TS: "T \<subseteq> affine hull S" | |
| 4553 | and S: "connected S" "a \<in> S" "b \<in> S" | |
| 4554 | obtains f g where "homeomorphism T T f g" "f a = b" | |
| 4555 |                     "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
 | |
| 4556 |                     "bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4557 | proof - | |
| 4558 | have 1: "\<exists>h k. homeomorphism T T h k \<and> h (f d) = d \<and> | |
| 4559 |               {x. \<not> (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. \<not> (h x = x \<and> k x = x)}"
 | |
| 4560 | if "d \<in> S" "f d \<in> S" and homfg: "homeomorphism T T f g" | |
| 4561 |         and S: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
 | |
| 4562 |         and bo: "bounded {x. \<not> (f x = x \<and> g x = x)}" for d f g
 | |
| 4563 | proof (intro exI conjI) | |
| 4564 | show homgf: "homeomorphism T T g f" | |
| 4565 | by (metis homeomorphism_symD homfg) | |
| 4566 | then show "g (f d) = d" | |
| 4567 | by (meson \<open>S \<subseteq> T\<close> homeomorphism_def subsetD \<open>d \<in> S\<close>) | |
| 4568 |     show "{x. \<not> (g x = x \<and> f x = x)} \<subseteq> S"
 | |
| 4569 | using S by blast | |
| 4570 |     show "bounded {x. \<not> (g x = x \<and> f x = x)}"
 | |
| 4571 | using bo by (simp add: conj_commute) | |
| 4572 | qed | |
| 4573 | have 2: "\<exists>f g. homeomorphism T T f g \<and> f x = f2 (f1 x) \<and> | |
| 4574 |                  {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4575 | if "x \<in> S" "f1 x \<in> S" "f2 (f1 x) \<in> S" | |
| 4576 | and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" | |
| 4577 |                 and sub: "{x. \<not> (f1 x = x \<and> g1 x = x)} \<subseteq> S"   "{x. \<not> (f2 x = x \<and> g2 x = x)} \<subseteq> S"
 | |
| 4578 |                 and bo: "bounded {x. \<not> (f1 x = x \<and> g1 x = x)}"  "bounded {x. \<not> (f2 x = x \<and> g2 x = x)}"
 | |
| 4579 | for x f1 f2 g1 g2 | |
| 4580 | proof (intro exI conjI) | |
| 4581 | show homgf: "homeomorphism T T (f2 \<circ> f1) (g1 \<circ> g2)" | |
| 4582 | by (metis homeomorphism_compose hom) | |
| 4583 | then show "(f2 \<circ> f1) x = f2 (f1 x)" | |
| 4584 | by force | |
| 4585 |     show "{x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)} \<subseteq> S"
 | |
| 4586 | using sub by force | |
| 4587 |     have "bounded ({x. \<not>(f1 x = x \<and> g1 x = x)} \<union> {x. \<not>(f2 x = x \<and> g2 x = x)})"
 | |
| 4588 | using bo by simp | |
| 4589 |     then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}"
 | |
| 4590 | by (rule bounded_subset) auto | |
| 4591 | qed | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4592 | have 3: "\<exists>U. openin (top_of_set S) U \<and> | 
| 69620 | 4593 | d \<in> U \<and> | 
| 4594 | (\<forall>x\<in>U. | |
| 4595 | \<exists>f g. homeomorphism T T f g \<and> f d = x \<and> | |
| 4596 |                         {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and>
 | |
| 4597 |                         bounded {x. \<not> (f x = x \<and> g x = x)})"
 | |
| 4598 | if "d \<in> S" for d | |
| 4599 | proof - | |
| 4600 | obtain r where "r > 0" and r: "ball d r \<inter> affine hull S \<subseteq> S" | |
| 4601 | by (metis \<open>d \<in> S\<close> ope openin_contains_ball) | |
| 4602 | have *: "\<exists>f g. homeomorphism T T f g \<and> f d = e \<and> | |
| 4603 |                    {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and>
 | |
| 4604 |                    bounded {x. \<not> (f x = x \<and> g x = x)}" if "e \<in> S" "e \<in> ball d r" for e
 | |
| 4605 | apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) | |
| 72372 | 4606 | using r \<open>S \<subseteq> T\<close> TS that | 
| 69620 | 4607 | apply (auto simp: \<open>d \<in> S\<close> \<open>0 < r\<close> hull_inc) | 
| 4608 | using bounded_subset by blast | |
| 4609 | show ?thesis | |
| 71769 | 4610 | by (rule_tac x="S \<inter> ball d r" in exI) (fastforce simp: openin_open_Int \<open>0 < r\<close> that intro: *) | 
| 69620 | 4611 | qed | 
| 4612 | have "\<exists>f g. homeomorphism T T f g \<and> f a = b \<and> | |
| 4613 |               {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 71769 | 4614 | by (rule connected_equivalence_relation [OF S]; blast intro: 1 2 3) | 
| 69620 | 4615 | then show ?thesis | 
| 4616 | using that by auto | |
| 4617 | qed | |
| 4618 | ||
| 4619 | ||
| 4620 | lemma homeomorphism_moving_points_exists_gen: | |
| 4621 | assumes K: "finite K" "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" | |
| 4622 | "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" | |
| 4623 | and "2 \<le> aff_dim S" | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4624 | and ope: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4625 | and "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" | 
| 4626 | shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> | |
| 4627 |                {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4628 | using assms | |
| 4629 | proof (induction K) | |
| 4630 | case empty | |
| 4631 | then show ?case | |
| 4632 | by (force simp: homeomorphism_ident) | |
| 4633 | next | |
| 4634 | case (insert i K) | |
| 4635 | then have xney: "\<And>j. \<lbrakk>j \<in> K; j \<noteq> i\<rbrakk> \<Longrightarrow> x i \<noteq> x j \<and> y i \<noteq> y j" | |
| 4636 | and pw: "pairwise (\<lambda>i j. x i \<noteq> x j \<and> y i \<noteq> y j) K" | |
| 4637 | and "x i \<in> S" "y i \<in> S" | |
| 4638 | and xyS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" | |
| 4639 | by (simp_all add: pairwise_insert) | |
| 4640 | obtain f g where homfg: "homeomorphism T T f g" and feq: "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" | |
| 4641 |                and fg_sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
 | |
| 4642 |                and bo_fg: "bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4643 | using insert.IH [OF xyS pw] insert.prems by (blast intro: that) | |
| 4644 | then have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> | |
| 4645 |                    {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4646 | using insert by blast | |
| 4647 | have aff_eq: "affine hull (S - y ` K) = affine hull S" | |
| 71769 | 4648 | proof (rule affine_hull_Diff [OF ope]) | 
| 4649 | show "finite (y ` K)" | |
| 4650 | by (simp add: insert.hyps(1)) | |
| 4651 | show "y ` K \<subset> S" | |
| 4652 | using \<open>y i \<in> S\<close> insert.hyps(2) xney xyS by fastforce | |
| 4653 | qed | |
| 69620 | 4654 | have f_in_S: "f x \<in> S" if "x \<in> S" for x | 
| 4655 | using homfg fg_sub homeomorphism_apply1 \<open>S \<subseteq> T\<close> | |
| 4656 | proof - | |
| 4657 | have "(f (f x) \<noteq> f x \<or> g (f x) \<noteq> f x) \<or> f x \<in> S" | |
| 4658 | by (metis \<open>S \<subseteq> T\<close> homfg subsetD homeomorphism_apply1 that) | |
| 4659 | then show ?thesis | |
| 4660 | using fg_sub by force | |
| 4661 | qed | |
| 4662 | obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" | |
| 4663 |                and hk_sub: "{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S - y ` K"
 | |
| 4664 |                and bo_hk:  "bounded {x. \<not> (h x = x \<and> k x = x)}"
 | |
| 4665 | proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4666 | show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" | 
| 69620 | 4667 | by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) | 
| 4668 | show "S - y ` K \<subseteq> T" | |
| 4669 | using \<open>S \<subseteq> T\<close> by auto | |
| 4670 | show "T \<subseteq> affine hull (S - y ` K)" | |
| 4671 | using insert by (simp add: aff_eq) | |
| 4672 | show "connected (S - y ` K)" | |
| 4673 | proof (rule connected_openin_diff_countable [OF \<open>connected S\<close> ope]) | |
| 4674 | show "\<not> collinear S" | |
| 4675 | using collinear_aff_dim \<open>2 \<le> aff_dim S\<close> by force | |
| 4676 | show "countable (y ` K)" | |
| 4677 | using countable_finite insert.hyps(1) by blast | |
| 4678 | qed | |
| 71769 | 4679 | have "\<And>k. \<lbrakk>f (x i) = y k; k \<in> K\<rbrakk> \<Longrightarrow> False" | 
| 69620 | 4680 | by (metis feq homfg \<open>x i \<in> S\<close> homeomorphism_def \<open>S \<subseteq> T\<close> \<open>i \<notin> K\<close> subsetCE xney xyS) | 
| 71769 | 4681 | then show "f (x i) \<in> S - y ` K" | 
| 4682 | by (auto simp: f_in_S \<open>x i \<in> S\<close>) | |
| 69620 | 4683 | show "y i \<in> S - y ` K" | 
| 4684 | using insert.hyps xney by (auto simp: \<open>y i \<in> S\<close>) | |
| 4685 | qed blast | |
| 4686 | show ?case | |
| 4687 | proof (intro exI conjI) | |
| 4688 | show "homeomorphism T T (h \<circ> f) (g \<circ> k)" | |
| 4689 | using homfg homhk homeomorphism_compose by blast | |
| 4690 | show "\<forall>i \<in> insert i K. (h \<circ> f) (x i) = y i" | |
| 4691 | using feq hk_sub by (auto simp: heq) | |
| 4692 |     show "{x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)} \<subseteq> S"
 | |
| 4693 | using fg_sub hk_sub by force | |
| 4694 |     have "bounded ({x. \<not>(f x = x \<and> g x = x)} \<union> {x. \<not>(h x = x \<and> k x = x)})"
 | |
| 4695 | using bo_fg bo_hk bounded_Un by blast | |
| 4696 |     then show "bounded {x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)}"
 | |
| 4697 | by (rule bounded_subset) auto | |
| 4698 | qed | |
| 4699 | qed | |
| 4700 | ||
| 70136 | 4701 | proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_points_exists: | 
| 69620 | 4702 | fixes S :: "'a::euclidean_space set" | 
| 4703 |   assumes 2: "2 \<le> DIM('a)" "open S" "connected S" "S \<subseteq> T" "finite K"
 | |
| 4704 | and KS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" | |
| 4705 | and pw: "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" | |
| 4706 | and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" | |
| 4707 | obtains f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" | |
| 4708 |                     "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (\<not> (f x = x \<and> g x = x))}"
 | |
| 4709 | proof (cases "S = {}")
 | |
| 4710 | case True | |
| 4711 | then show ?thesis | |
| 4712 | using KS homeomorphism_ident that by fastforce | |
| 4713 | next | |
| 4714 | case False | |
| 4715 | then have affS: "affine hull S = UNIV" | |
| 4716 | by (simp add: affine_hull_open \<open>open S\<close>) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4717 | then have ope: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 4718 | using \<open>open S\<close> open_openin by auto | 
| 4719 |   have "2 \<le> DIM('a)" by (rule 2)
 | |
| 4720 | also have "\<dots> = aff_dim (UNIV :: 'a set)" | |
| 4721 | by simp | |
| 4722 | also have "\<dots> \<le> aff_dim S" | |
| 4723 | by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) | |
| 4724 | finally have "2 \<le> aff_dim S" | |
| 4725 | by linarith | |
| 4726 | then show ?thesis | |
| 4727 | using homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> KS pw _ ope S] that by fastforce | |
| 4728 | qed | |
| 4729 | ||
| 70136 | 4730 | subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close> | 
| 69620 | 4731 | |
| 4732 | lemma homeomorphism_grouping_point_1: | |
| 4733 | fixes a::real and c::real | |
| 4734 | assumes "a < b" "c < d" | |
| 4735 | obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" | |
| 4736 | proof - | |
| 4737 | define f where "f \<equiv> \<lambda>x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" | |
| 4738 | have "\<exists>g. homeomorphism (cbox a b) (cbox c d) f g" | |
| 4739 | proof (rule homeomorphism_compact) | |
| 4740 | show "continuous_on (cbox a b) f" | |
| 71769 | 4741 | unfolding f_def by (intro continuous_intros) | 
| 69620 | 4742 |     have "f ` {a..b} = {c..d}"
 | 
| 4743 | unfolding f_def image_affinity_atLeastAtMost | |
| 71172 | 4744 | using assms sum_sqs_eq by (auto simp: field_split_simps) | 
| 69620 | 4745 | then show "f ` cbox a b = cbox c d" | 
| 4746 | by auto | |
| 4747 | show "inj_on f (cbox a b)" | |
| 4748 | unfolding f_def inj_on_def using assms by auto | |
| 4749 | qed auto | |
| 4750 | then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. | |
| 4751 | then show ?thesis | |
| 4752 | proof | |
| 4753 | show "f a = c" | |
| 4754 | by (simp add: f_def) | |
| 4755 | show "f b = d" | |
| 71172 | 4756 | using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps) | 
| 69620 | 4757 | qed | 
| 4758 | qed | |
| 4759 | ||
| 4760 | lemma homeomorphism_grouping_point_2: | |
| 4761 | fixes a::real and w::real | |
| 4762 | assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" | |
| 4763 | and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" | |
| 4764 | and "b \<in> cbox a c" "v \<in> cbox u w" | |
| 4765 | and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" | |
| 4766 | obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" | |
| 4767 | "\<And>x. x \<in> cbox a b \<Longrightarrow> f x = f1 x" "\<And>x. x \<in> cbox b c \<Longrightarrow> f x = f2 x" | |
| 4768 | proof - | |
| 4769 | have le: "a \<le> b" "b \<le> c" "u \<le> v" "v \<le> w" | |
| 4770 | using assms by simp_all | |
| 4771 | then have ac: "cbox a c = cbox a b \<union> cbox b c" and uw: "cbox u w = cbox u v \<union> cbox v w" | |
| 4772 | by auto | |
| 4773 | define f where "f \<equiv> \<lambda>x. if x \<le> b then f1 x else f2 x" | |
| 4774 | have "\<exists>g. homeomorphism (cbox a c) (cbox u w) f g" | |
| 4775 | proof (rule homeomorphism_compact) | |
| 4776 | have cf1: "continuous_on (cbox a b) f1" | |
| 4777 | using hom_ab homeomorphism_cont1 by blast | |
| 4778 | have cf2: "continuous_on (cbox b c) f2" | |
| 4779 | using hom_bc homeomorphism_cont1 by blast | |
| 4780 | show "continuous_on (cbox a c) f" | |
| 72372 | 4781 | unfolding f_def using le eq | 
| 4782 | by (force intro: continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) | |
| 69620 | 4783 | have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" | 
| 4784 | unfolding f_def using eq by force+ | |
| 4785 | then show "f ` cbox a c = cbox u w" | |
| 71769 | 4786 | unfolding ac uw image_Un by (metis hom_ab hom_bc homeomorphism_def) | 
| 69620 | 4787 | have neq12: "f1 x \<noteq> f2 y" if x: "a \<le> x" "x \<le> b" and y: "b < y" "y \<le> c" for x y | 
| 4788 | proof - | |
| 4789 | have "f1 x \<in> cbox u v" | |
| 4790 | by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) | |
| 4791 | moreover have "f2 y \<in> cbox v w" | |
| 4792 | by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) | |
| 4793 | moreover have "f2 y \<noteq> f2 b" | |
| 4794 | by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) | |
| 4795 | ultimately show ?thesis | |
| 4796 | using le eq by simp | |
| 4797 | qed | |
| 4798 | have "inj_on f1 (cbox a b)" | |
| 4799 | by (metis (full_types) hom_ab homeomorphism_def inj_onI) | |
| 4800 | moreover have "inj_on f2 (cbox b c)" | |
| 4801 | by (metis (full_types) hom_bc homeomorphism_def inj_onI) | |
| 4802 | ultimately show "inj_on f (cbox a c)" | |
| 4803 | apply (simp (no_asm) add: inj_on_def) | |
| 4804 | apply (simp add: f_def inj_on_eq_iff) | |
| 71769 | 4805 | using neq12 by force | 
| 69620 | 4806 | qed auto | 
| 4807 | then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. | |
| 4808 | then show ?thesis | |
| 71769 | 4809 | using eq f_def le that by force | 
| 69620 | 4810 | qed | 
| 4811 | ||
| 4812 | lemma homeomorphism_grouping_point_3: | |
| 4813 | fixes a::real | |
| 4814 | assumes cbox_sub: "cbox c d \<subseteq> box a b" "cbox u v \<subseteq> box a b" | |
| 4815 |       and box_ne: "box c d \<noteq> {}" "box u v \<noteq> {}"
 | |
| 4816 | obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" | |
| 4817 | "\<And>x. x \<in> cbox c d \<Longrightarrow> f x \<in> cbox u v" | |
| 4818 | proof - | |
| 4819 |   have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \<noteq> {}"
 | |
| 4820 | using assms | |
| 4821 | by (simp_all add: cbox_sub subset_eq) | |
| 4822 | obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" | |
| 4823 | and f1_eq: "f1 a = a" "f1 c = u" | |
| 4824 | using homeomorphism_grouping_point_1 [OF \<open>a < c\<close> \<open>a < u\<close>] . | |
| 4825 | obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" | |
| 4826 | and f2_eq: "f2 c = u" "f2 d = v" | |
| 4827 | using homeomorphism_grouping_point_1 [OF \<open>c < d\<close> \<open>u < v\<close>] . | |
| 4828 | obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" | |
| 4829 | and f3_eq: "f3 d = v" "f3 b = b" | |
| 4830 | using homeomorphism_grouping_point_1 [OF \<open>d < b\<close> \<open>v < b\<close>] . | |
| 4831 | obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" | |
| 4832 | and f4_eq: "\<And>x. x \<in> cbox a c \<Longrightarrow> f4 x = f1 x" "\<And>x. x \<in> cbox c d \<Longrightarrow> f4 x = f2 x" | |
| 4833 | using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) | |
| 4834 | obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" | |
| 4835 | and f_eq: "\<And>x. x \<in> cbox a d \<Longrightarrow> f x = f4 x" "\<And>x. x \<in> cbox d b \<Longrightarrow> f x = f3 x" | |
| 4836 | using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) | |
| 4837 | show ?thesis | |
| 71769 | 4838 | proof (rule that [OF fg]) | 
| 4839 | show "f x \<in> cbox u v" if "x \<in> cbox c d" for x | |
| 4840 | using that f4_eq f_eq homeomorphism_image1 [OF 2] | |
| 4841 | by (metis atLeastAtMost_iff box_real(2) image_eqI less(1) less_eq_real_def order_trans) | |
| 4842 | qed | |
| 69620 | 4843 | qed | 
| 4844 | ||
| 4845 | ||
| 4846 | lemma homeomorphism_grouping_point_4: | |
| 4847 | fixes T :: "real set" | |
| 4848 |   assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
 | |
| 4849 | obtains f g where "homeomorphism T T f g" | |
| 4850 |                     "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
 | |
| 4851 |                     "bounded {x. (\<not> (f x = x \<and> g x = x))}"
 | |
| 4852 | proof - | |
| 4853 |   obtain c d where "box c d \<noteq> {}" "cbox c d \<subseteq> U"
 | |
| 4854 | proof - | |
| 4855 | obtain u where "u \<in> U" | |
| 4856 |       using \<open>U \<noteq> {}\<close> by blast
 | |
| 4857 | then obtain e where "e > 0" "cball u e \<subseteq> U" | |
| 4858 | using \<open>open U\<close> open_contains_cball by blast | |
| 4859 | then show ?thesis | |
| 4860 | by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) | |
| 4861 | qed | |
| 4862 | have "compact K" | |
| 4863 | by (simp add: \<open>finite K\<close> finite_imp_compact) | |
| 4864 |   obtain a b where "box a b \<noteq> {}" "K \<subseteq> cbox a b" "cbox a b \<subseteq> S"
 | |
| 4865 |   proof (cases "K = {}")
 | |
| 4866 | case True then show ?thesis | |
| 4867 |       using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> that by blast
 | |
| 4868 | next | |
| 4869 | case False | |
| 4870 | then obtain a b where "a \<in> K" "b \<in> K" | |
| 4871 | and a: "\<And>x. x \<in> K \<Longrightarrow> a \<le> x" and b: "\<And>x. x \<in> K \<Longrightarrow> x \<le> b" | |
| 4872 | using compact_attains_inf compact_attains_sup by (metis \<open>compact K\<close>)+ | |
| 4873 | obtain e where "e > 0" "cball b e \<subseteq> S" | |
| 4874 | using \<open>open S\<close> open_contains_cball | |
| 4875 | by (metis \<open>b \<in> K\<close> \<open>K \<subseteq> S\<close> subsetD) | |
| 4876 | show ?thesis | |
| 4877 | proof | |
| 4878 |       show "box a (b + e) \<noteq> {}"
 | |
| 4879 | using \<open>0 < e\<close> \<open>b \<in> K\<close> a by force | |
| 4880 | show "K \<subseteq> cbox a (b + e)" | |
| 4881 | using \<open>0 < e\<close> a b by fastforce | |
| 4882 | have "a \<in> S" | |
| 4883 | using \<open>a \<in> K\<close> assms(6) by blast | |
| 4884 | have "b + e \<in> S" | |
| 4885 | using \<open>0 < e\<close> \<open>cball b e \<subseteq> S\<close> by (force simp: dist_norm) | |
| 4886 | show "cbox a (b + e) \<subseteq> S" | |
| 4887 | using \<open>a \<in> S\<close> \<open>b + e \<in> S\<close> \<open>connected S\<close> connected_contains_Icc by auto | |
| 4888 | qed | |
| 4889 | qed | |
| 4890 | obtain w z where "cbox w z \<subseteq> S" and sub_wz: "cbox a b \<union> cbox c d \<subseteq> box w z" | |
| 4891 | proof - | |
| 4892 | have "a \<in> S" "b \<in> S" | |
| 4893 |       using \<open>box a b \<noteq> {}\<close> \<open>cbox a b \<subseteq> S\<close> by auto
 | |
| 4894 | moreover have "c \<in> S" "d \<in> S" | |
| 4895 |       using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> by force+
 | |
| 4896 | ultimately have "min a c \<in> S" "max b d \<in> S" | |
| 4897 | by linarith+ | |
| 4898 | then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \<subseteq> S" "e2 > 0" "cball (max b d) e2 \<subseteq> S" | |
| 4899 | using \<open>open S\<close> open_contains_cball by metis | |
| 4900 | then have *: "min a c - e1 \<in> S" "max b d + e2 \<in> S" | |
| 4901 | by (auto simp: dist_norm) | |
| 4902 | show ?thesis | |
| 4903 | proof | |
| 4904 | show "cbox (min a c - e1) (max b d+ e2) \<subseteq> S" | |
| 4905 | using * \<open>connected S\<close> connected_contains_Icc by auto | |
| 4906 | show "cbox a b \<union> cbox c d \<subseteq> box (min a c - e1) (max b d + e2)" | |
| 4907 | using \<open>0 < e1\<close> \<open>0 < e2\<close> by auto | |
| 4908 | qed | |
| 4909 | qed | |
| 4910 | then | |
| 4911 | obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" | |
| 4912 | and "f w = w" "f z = z" | |
| 4913 | and fin: "\<And>x. x \<in> cbox a b \<Longrightarrow> f x \<in> cbox c d" | |
| 4914 | using homeomorphism_grouping_point_3 [of a b w z c d] | |
| 4915 |     using \<open>box a b \<noteq> {}\<close> \<open>box c d \<noteq> {}\<close> by blast
 | |
| 4916 | have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" | |
| 4917 | using hom homeomorphism_def by blast+ | |
| 4918 | define f' where "f' \<equiv> \<lambda>x. if x \<in> cbox w z then f x else x" | |
| 4919 | define g' where "g' \<equiv> \<lambda>x. if x \<in> cbox w z then g x else x" | |
| 4920 | show ?thesis | |
| 4921 | proof | |
| 4922 | have T: "cbox w z \<union> (T - box w z) = T" | |
| 4923 | using \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> by auto | |
| 4924 | show "homeomorphism T T f' g'" | |
| 4925 | proof | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 4926 | have clo: "closedin (top_of_set (cbox w z \<union> (T - box w z))) (T - box w z)" | 
| 69620 | 4927 | by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) | 
| 72372 | 4928 | have "\<And>x. \<lbrakk>w \<le> x \<and> x \<le> z; w < x \<longrightarrow> \<not> x < z\<rbrakk> \<Longrightarrow> f x = x" | 
| 4929 | using \<open>f w = w\<close> \<open>f z = z\<close> by auto | |
| 4930 | moreover have "\<And>x. \<lbrakk>w \<le> x \<and> x \<le> z; w < x \<longrightarrow> \<not> x < z\<rbrakk> \<Longrightarrow> g x = x" | |
| 4931 | using \<open>f w = w\<close> \<open>f z = z\<close> hom homeomorphism_apply1 by fastforce | |
| 4932 | ultimately | |
| 69620 | 4933 | have "continuous_on (cbox w z \<union> (T - box w z)) f'" "continuous_on (cbox w z \<union> (T - box w z)) g'" | 
| 4934 | unfolding f'_def g'_def | |
| 72372 | 4935 | by (intro continuous_on_cases_local contfg continuous_on_id clo; auto simp: closed_subset)+ | 
| 69620 | 4936 | then show "continuous_on T f'" "continuous_on T g'" | 
| 4937 | by (simp_all only: T) | |
| 4938 | show "f' ` T \<subseteq> T" | |
| 4939 | unfolding f'_def | |
| 4940 | by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) | |
| 4941 | show "g' ` T \<subseteq> T" | |
| 4942 | unfolding g'_def | |
| 4943 | by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) | |
| 4944 | show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" | |
| 4945 | unfolding f'_def g'_def | |
| 4946 | using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce | |
| 4947 | show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" | |
| 4948 | unfolding f'_def g'_def | |
| 4949 | using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce | |
| 4950 | qed | |
| 4951 | show "\<And>x. x \<in> K \<Longrightarrow> f' x \<in> U" | |
| 4952 | using fin sub_wz \<open>K \<subseteq> cbox a b\<close> \<open>cbox c d \<subseteq> U\<close> by (force simp: f'_def) | |
| 4953 |     show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S"
 | |
| 4954 | using \<open>cbox w z \<subseteq> S\<close> by (auto simp: f'_def g'_def) | |
| 4955 |     show "bounded {x. \<not> (f' x = x \<and> g' x = x)}"
 | |
| 71769 | 4956 | proof (rule bounded_subset [of "cbox w z"]) | 
| 4957 | show "bounded (cbox w z)" | |
| 4958 | using bounded_cbox by blast | |
| 4959 |       show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> cbox w z"
 | |
| 4960 | by (auto simp: f'_def g'_def) | |
| 4961 | qed | |
| 69620 | 4962 | qed | 
| 4963 | qed | |
| 4964 | ||
| 70136 | 4965 | proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists: | 
| 69620 | 4966 | fixes S :: "'a::euclidean_space set" | 
| 4967 |   assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
 | |
| 4968 |   obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
 | |
| 4969 |                     "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
 | |
| 4970 | proof (cases "2 \<le> DIM('a)")
 | |
| 4971 | case True | |
| 4972 | have TS: "T \<subseteq> affine hull S" | |
| 4973 | using affine_hull_open assms by blast | |
| 4974 | have "infinite U" | |
| 4975 |     using \<open>open U\<close> \<open>U \<noteq> {}\<close> finite_imp_not_open by blast
 | |
| 4976 | then obtain P where "P \<subseteq> U" "finite P" "card K = card P" | |
| 4977 | using infinite_arbitrarily_large by metis | |
| 4978 | then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" | |
| 4979 | using \<open>finite K\<close> finite_same_card_bij by blast | |
| 4980 |   obtain f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f (id i) = \<gamma> i" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 4981 | proof (rule homeomorphism_moving_points_exists [OF True \<open>open S\<close> \<open>connected S\<close> \<open>S \<subseteq> T\<close> \<open>finite K\<close>]) | |
| 4982 | show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" | |
| 4983 | using \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> \<open>U \<subseteq> S\<close> bij_betwE by blast | |
| 4984 | show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" | |
| 4985 | using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) | |
| 4986 | qed (use affine_hull_open assms that in auto) | |
| 4987 | then show ?thesis | |
| 78474 | 4988 | using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that) | 
| 69620 | 4989 | next | 
| 4990 | case False | |
| 4991 |   with DIM_positive have "DIM('a) = 1"
 | |
| 4992 | by (simp add: dual_order.antisym) | |
| 4993 | then obtain h::"'a \<Rightarrow>real" and j | |
| 4994 | where "linear h" "linear j" | |
| 4995 | and noh: "\<And>x. norm(h x) = norm x" and noj: "\<And>y. norm(j y) = norm y" | |
| 4996 | and hj: "\<And>x. j(h x) = x" "\<And>y. h(j y) = y" | |
| 4997 | and ranh: "surj h" | |
| 4998 | using isomorphisms_UNIV_UNIV | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72372diff
changeset | 4999 | by (metis (mono_tags, opaque_lifting) DIM_real UNIV_eq_I range_eqI) | 
| 69620 | 5000 | obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" | 
| 5001 | and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" | |
| 5002 |                and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S"
 | |
| 5003 |                and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 5004 | apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) | |
| 5005 | by (simp_all add: assms image_mono \<open>linear h\<close> open_surjective_linear_image connected_linear_image ranh) | |
| 5006 | have jf: "j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" for x | |
| 5007 | by (metis hj) | |
| 5008 | have jg: "j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" for x | |
| 5009 | by (metis hj) | |
| 5010 | have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y | |
| 5011 | by (simp_all add: \<open>linear h\<close> \<open>linear j\<close> linear_linear linear_continuous_on) | |
| 5012 | show ?thesis | |
| 5013 | proof | |
| 5014 | show "homeomorphism T T (j \<circ> f \<circ> h) (j \<circ> g \<circ> h)" | |
| 5015 | proof | |
| 5016 | show "continuous_on T (j \<circ> f \<circ> h)" "continuous_on T (j \<circ> g \<circ> h)" | |
| 5017 | using hom homeomorphism_def | |
| 5018 | by (blast intro: continuous_on_compose cont_hj)+ | |
| 5019 | show "(j \<circ> f \<circ> h) ` T \<subseteq> T" "(j \<circ> g \<circ> h) ` T \<subseteq> T" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72372diff
changeset | 5020 | by auto (metis (mono_tags, opaque_lifting) hj(1) hom homeomorphism_def imageE imageI)+ | 
| 69620 | 5021 | show "\<And>x. x \<in> T \<Longrightarrow> (j \<circ> g \<circ> h) ((j \<circ> f \<circ> h) x) = x" | 
| 5022 | using hj hom homeomorphism_apply1 by fastforce | |
| 5023 | show "\<And>y. y \<in> T \<Longrightarrow> (j \<circ> f \<circ> h) ((j \<circ> g \<circ> h) y) = y" | |
| 5024 | using hj hom homeomorphism_apply2 by fastforce | |
| 5025 | qed | |
| 5026 |     show "{x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} \<subseteq> S"
 | |
| 72372 | 5027 | proof (clarsimp simp: jf jg hj) | 
| 5028 | show "f (h x) = h x \<longrightarrow> g (h x) \<noteq> h x \<Longrightarrow> x \<in> S" for x | |
| 5029 | using sub [THEN subsetD, of "h x"] hj by simp (metis imageE) | |
| 5030 | qed | |
| 69620 | 5031 |     have "bounded (j ` {x. (\<not> (f x = x \<and> g x = x))})"
 | 
| 5032 | by (rule bounded_linear_image [OF bou]) (use \<open>linear j\<close> linear_conv_bounded_linear in auto) | |
| 5033 | moreover | |
| 5034 |     have *: "{x. \<not>((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
 | |
| 5035 | using hj by (auto simp: jf jg image_iff, metis+) | |
| 5036 |     ultimately show "bounded {x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)}"
 | |
| 5037 | by metis | |
| 5038 | show "\<And>x. x \<in> K \<Longrightarrow> (j \<circ> f \<circ> h) x \<in> U" | |
| 5039 | using f hj by fastforce | |
| 5040 | qed | |
| 5041 | qed | |
| 5042 | ||
| 5043 | ||
| 70136 | 5044 | proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists_gen: | 
| 69620 | 5045 | fixes S :: "'a::euclidean_space set" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 5046 | assumes opeU: "openin (top_of_set S) U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 5047 | and opeS: "openin (top_of_set (affine hull S)) S" | 
| 69620 | 5048 |       and "U \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
 | 
| 5049 |   obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
 | |
| 5050 |                     "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
 | |
| 5051 | proof (cases "2 \<le> aff_dim S") | |
| 5052 | case True | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 5053 | have opeU': "openin (top_of_set (affine hull S)) U" | 
| 69620 | 5054 | using opeS opeU openin_trans by blast | 
| 5055 | obtain u where "u \<in> U" "u \<in> S" | |
| 5056 |     using \<open>U \<noteq> {}\<close> opeU openin_imp_subset by fastforce+
 | |
| 5057 | have "infinite U" | |
| 72372 | 5058 | proof (rule infinite_openin [OF opeU \<open>u \<in> U\<close>]) | 
| 5059 | show "u islimpt S" | |
| 5060 | using True \<open>u \<in> S\<close> assms(8) connected_imp_perfect_aff_dim by fastforce | |
| 5061 | qed | |
| 69620 | 5062 | then obtain P where "P \<subseteq> U" "finite P" "card K = card P" | 
| 5063 | using infinite_arbitrarily_large by metis | |
| 5064 | then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" | |
| 5065 | using \<open>finite K\<close> finite_same_card_bij by blast | |
| 5066 | have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(id i) = \<gamma> i) \<and> | |
| 5067 |                {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 5068 | proof (rule homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> _ _ True opeS S]) | |
| 5069 | show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" | |
| 5070 | by (metis id_apply opeU openin_contains_cball subsetCE \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> bij_betwE) | |
| 5071 | show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" | |
| 5072 | using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) | |
| 5073 | qed | |
| 5074 | then show ?thesis | |
| 78474 | 5075 | using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that) | 
| 69620 | 5076 | next | 
| 5077 | case False | |
| 5078 | with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith | |
| 5079 | then show ?thesis | |
| 5080 | proof cases | |
| 5081 | assume "aff_dim S = -1" | |
| 5082 |     then have "S = {}"
 | |
| 5083 | using aff_dim_empty by blast | |
| 5084 | then have "False" | |
| 5085 |       using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast
 | |
| 5086 | then show ?thesis .. | |
| 5087 | next | |
| 5088 | assume "aff_dim S = 0" | |
| 5089 |     then obtain a where "S = {a}"
 | |
| 5090 | using aff_dim_eq_0 by blast | |
| 5091 | then have "K \<subseteq> U" | |
| 5092 |       using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast
 | |
| 5093 | show ?thesis | |
| 72372 | 5094 | using \<open>K \<subseteq> U\<close> by (intro that [of id id]) (auto intro: homeomorphismI) | 
| 69620 | 5095 | next | 
| 5096 | assume "aff_dim S = 1" | |
| 5097 | then have "affine hull S homeomorphic (UNIV :: real set)" | |
| 5098 | by (auto simp: homeomorphic_affine_sets) | |
| 5099 | then obtain h::"'a\<Rightarrow>real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" | |
| 5100 | using homeomorphic_def by blast | |
| 5101 | then have h: "\<And>x. x \<in> affine hull S \<Longrightarrow> j(h(x)) = x" and j: "\<And>y. j y \<in> affine hull S \<and> h(j y) = y" | |
| 5102 | by (auto simp: homeomorphism_def) | |
| 5103 | have connh: "connected (h ` S)" | |
| 5104 | by (meson Topological_Spaces.connected_continuous_image \<open>connected S\<close> homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) | |
| 5105 | have hUS: "h ` U \<subseteq> h ` S" | |
| 5106 | by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 5107 | have opn: "openin (top_of_set (affine hull S)) U \<Longrightarrow> open (h ` U)" for U | 
| 69620 | 5108 | using homeomorphism_imp_open_map [OF homhj] by simp | 
| 5109 | have "open (h ` U)" "open (h ` S)" | |
| 5110 | by (auto intro: opeS opeU openin_trans opn) | |
| 5111 | then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" | |
| 5112 | and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" | |
| 5113 |                  and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S"
 | |
| 5114 |                  and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}"
 | |
| 5115 | apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) | |
| 5116 | using assms by (auto simp: connh hUS) | |
| 5117 | have jf: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" | |
| 5118 | by (metis h j) | |
| 5119 | have jg: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" | |
| 5120 | by (metis h j) | |
| 5121 | have cont_hj: "continuous_on T h" "continuous_on Y j" for Y | |
| 72372 | 5122 | proof (rule continuous_on_subset [OF _ \<open>T \<subseteq> affine hull S\<close>]) | 
| 5123 | show "continuous_on (affine hull S) h" | |
| 5124 | using homeomorphism_def homhj by blast | |
| 5125 | qed (meson continuous_on_subset homeomorphism_def homhj top_greatest) | |
| 69620 | 5126 | define f' where "f' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> f \<circ> h) x else x" | 
| 5127 | define g' where "g' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> g \<circ> h) x else x" | |
| 5128 | show ?thesis | |
| 5129 | proof | |
| 5130 | show "homeomorphism T T f' g'" | |
| 5131 | proof | |
| 5132 | have "continuous_on T (j \<circ> f \<circ> h)" | |
| 72372 | 5133 | using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast | 
| 69620 | 5134 | then show "continuous_on T f'" | 
| 5135 | apply (rule continuous_on_eq) | |
| 5136 | using \<open>T \<subseteq> affine hull S\<close> f'_def by auto | |
| 5137 | have "continuous_on T (j \<circ> g \<circ> h)" | |
| 72372 | 5138 | using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast | 
| 69620 | 5139 | then show "continuous_on T g'" | 
| 5140 | apply (rule continuous_on_eq) | |
| 5141 | using \<open>T \<subseteq> affine hull S\<close> g'_def by auto | |
| 5142 | show "f' ` T \<subseteq> T" | |
| 5143 | proof (clarsimp simp: f'_def) | |
| 5144 | fix x assume "x \<in> T" | |
| 5145 | then have "f (h x) \<in> h ` T" | |
| 5146 | by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) | |
| 5147 | then show "j (f (h x)) \<in> T" | |
| 5148 | using \<open>T \<subseteq> affine hull S\<close> h by auto | |
| 5149 | qed | |
| 5150 | show "g' ` T \<subseteq> T" | |
| 5151 | proof (clarsimp simp: g'_def) | |
| 5152 | fix x assume "x \<in> T" | |
| 5153 | then have "g (h x) \<in> h ` T" | |
| 5154 | by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) | |
| 5155 | then show "j (g (h x)) \<in> T" | |
| 5156 | using \<open>T \<subseteq> affine hull S\<close> h by auto | |
| 5157 | qed | |
| 5158 | show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" | |
| 78474 | 5159 | using h j hom homeomorphism_apply1 by (fastforce simp: f'_def g'_def) | 
| 69620 | 5160 | show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" | 
| 78474 | 5161 | using h j hom homeomorphism_apply2 by (fastforce simp: f'_def g'_def) | 
| 69620 | 5162 | qed | 
| 5163 | next | |
| 72372 | 5164 | have \<section>: "\<And>x y. \<lbrakk>x \<in> affine hull S; h x = h y; y \<in> S\<rbrakk> \<Longrightarrow> x \<in> S" | 
| 5165 | by (metis h hull_inc) | |
| 69620 | 5166 |       show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S"
 | 
| 72372 | 5167 | using sub by (simp add: f'_def g'_def jf jg) (force elim: \<section>) | 
| 69620 | 5168 | next | 
| 5169 |       have "compact (j ` closure {x. \<not> (f x = x \<and> g x = x)})"
 | |
| 5170 | using bou by (auto simp: compact_continuous_image cont_hj) | |
| 5171 |       then have "bounded (j ` {x. \<not> (f x = x \<and> g x = x)})"
 | |
| 5172 | by (rule bounded_closure_image [OF compact_imp_bounded]) | |
| 5173 | moreover | |
| 5174 |       have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
 | |
| 5175 | using h j by (auto simp: image_iff; metis) | |
| 5176 |       ultimately have "bounded {x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x}"
 | |
| 5177 | by metis | |
| 5178 |       then show "bounded {x. \<not> (f' x = x \<and> g' x = x)}"
 | |
| 5179 | by (simp add: f'_def g'_def Collect_mono bounded_subset) | |
| 5180 | next | |
| 5181 | show "f' x \<in> U" if "x \<in> K" for x | |
| 5182 | proof - | |
| 5183 | have "U \<subseteq> S" | |
| 5184 | using opeU openin_imp_subset by blast | |
| 5185 | then have "j (f (h x)) \<in> U" | |
| 5186 | using f h hull_subset that by fastforce | |
| 5187 | then show "f' x \<in> U" | |
| 5188 | using \<open>K \<subseteq> S\<close> S f'_def that by auto | |
| 5189 | qed | |
| 5190 | qed | |
| 5191 | qed | |
| 5192 | qed | |
| 5193 | ||
| 5194 | ||
| 5195 | subsection\<open>Nullhomotopic mappings\<close> | |
| 5196 | ||
| 5197 | text\<open> A mapping out of a sphere is nullhomotopic iff it extends to the ball. | |
| 5198 | This even works out in the degenerate cases when the radius is \<open>\<le>\<close> 0, and | |
| 5199 | we also don't need to explicitly assume continuity since it's already implicit | |
| 5200 | in both sides of the equivalence.\<close> | |
| 5201 | ||
| 5202 | lemma nullhomotopic_from_lemma: | |
| 5203 |   assumes contg: "continuous_on (cball a r - {a}) g"
 | |
| 5204 | and fa: "\<And>e. 0 < e | |
| 5205 | \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>x. x \<noteq> a \<and> norm(x - a) < d \<longrightarrow> norm(g x - f a) < e)" | |
| 5206 | and r: "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> f x = g x" | |
| 5207 | shows "continuous_on (cball a r) f" | |
| 5208 | proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) | |
| 5209 | fix x | |
| 5210 | assume x: "dist a x \<le> r" | |
| 5211 | show "continuous (at x within cball a r) f" | |
| 5212 | proof (cases "x=a") | |
| 5213 | case True | |
| 5214 | then show ?thesis | |
| 5215 | by (metis continuous_within_eps_delta fa dist_norm dist_self r) | |
| 5216 | next | |
| 5217 | case False | |
| 5218 | show ?thesis | |
| 5219 | proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) | |
| 5220 | have "\<exists>d>0. \<forall>x'\<in>cball a r. | |
| 5221 | dist x' x < d \<longrightarrow> dist (g x') (g x) < e" if "e>0" for e | |
| 5222 | proof - | |
| 5223 | obtain d where "d > 0" | |
| 5224 | and d: "\<And>x'. \<lbrakk>dist x' a \<le> r; x' \<noteq> a; dist x' x < d\<rbrakk> \<Longrightarrow> | |
| 5225 | dist (g x') (g x) < e" | |
| 5226 | using contg False x \<open>e>0\<close> | |
| 78474 | 5227 | unfolding continuous_on_iff by (fastforce simp: dist_commute intro: that) | 
| 69620 | 5228 | show ?thesis | 
| 5229 | using \<open>d > 0\<close> \<open>x \<noteq> a\<close> | |
| 5230 | by (rule_tac x="min d (norm(x - a))" in exI) | |
| 5231 | (auto simp: dist_commute dist_norm [symmetric] intro!: d) | |
| 5232 | qed | |
| 5233 | then show "continuous (at x within cball a r) g" | |
| 5234 | using contg False by (auto simp: continuous_within_eps_delta) | |
| 5235 | show "0 < norm (x - a)" | |
| 5236 | using False by force | |
| 5237 | show "x \<in> cball a r" | |
| 5238 | by (simp add: x) | |
| 5239 | show "\<And>x'. \<lbrakk>x' \<in> cball a r; dist x' x < norm (x - a)\<rbrakk> | |
| 5240 | \<Longrightarrow> g x' = f x'" | |
| 5241 | by (metis dist_commute dist_norm less_le r) | |
| 5242 | qed | |
| 5243 | qed | |
| 5244 | qed | |
| 5245 | ||
| 5246 | proposition nullhomotopic_from_sphere_extension: | |
| 5247 | fixes f :: "'M::euclidean_space \<Rightarrow> 'a::real_normed_vector" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5248 | shows "(\<exists>c. homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)) \<longleftrightarrow> | 
| 69620 | 5249 | (\<exists>g. continuous_on (cball a r) g \<and> g ` (cball a r) \<subseteq> S \<and> | 
| 5250 | (\<forall>x \<in> sphere a r. g x = f x))" | |
| 5251 | (is "?lhs = ?rhs") | |
| 5252 | proof (cases r "0::real" rule: linorder_cases) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5253 | case less | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5254 | then show ?thesis | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5255 | by (simp add: homotopic_on_emptyI) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5256 | next | 
| 69620 | 5257 | case equal | 
| 71768 | 5258 | show ?thesis | 
| 5259 | proof | |
| 5260 | assume L: ?lhs | |
| 5261 | with equal have [simp]: "f a \<in> S" | |
| 5262 | using homotopic_with_imp_subset1 by fastforce | |
| 5263 | obtain h:: "real \<times> 'M \<Rightarrow> 'a" | |
| 5264 |       where h: "continuous_on ({0..1} \<times> {a}) h" "h ` ({0..1} \<times> {a}) \<subseteq> S" "h (0, a) = f a"    
 | |
| 5265 | using L equal by (auto simp: homotopic_with) | |
| 5266 | then have "continuous_on (cball a r) (\<lambda>x. h (0, a))" "(\<lambda>x. h (0, a)) ` cball a r \<subseteq> S" | |
| 5267 | by (auto simp: equal) | |
| 5268 | then show ?rhs | |
| 5269 | using h(3) local.equal by force | |
| 5270 | next | |
| 5271 | assume ?rhs | |
| 5272 | then show ?lhs | |
| 78474 | 5273 | using equal continuous_on_const by (force simp: homotopic_with) | 
| 71768 | 5274 | qed | 
| 69620 | 5275 | next | 
| 5276 | case greater | |
| 5277 |   let ?P = "continuous_on {x. norm(x - a) = r} f \<and> f ` {x. norm(x - a) = r} \<subseteq> S"
 | |
| 5278 | have ?P if ?lhs using that | |
| 5279 | proof | |
| 5280 | fix c | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5281 | assume c: "homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5282 | then have contf: "continuous_on (sphere a r) f" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5283 | by (metis homotopic_with_imp_continuous) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5284 | moreover have fim: "f ` sphere a r \<subseteq> S" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5285 | by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps) | 
| 69620 | 5286 | show ?P | 
| 5287 | using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) | |
| 5288 | qed | |
| 5289 | moreover have ?P if ?rhs using that | |
| 5290 | proof | |
| 5291 | fix g | |
| 5292 | assume g: "continuous_on (cball a r) g \<and> g ` cball a r \<subseteq> S \<and> (\<forall>xa\<in>sphere a r. g xa = f xa)" | |
| 71768 | 5293 |     then have "f ` {x. norm (x - a) = r} \<subseteq> S"
 | 
| 5294 | using sphere_cball [of a r] unfolding image_subset_iff sphere_def | |
| 5295 | by (metis dist_commute dist_norm mem_Collect_eq subset_eq) | |
| 5296 | with g show ?P | |
| 5297 | by (auto simp: dist_norm norm_minus_commute elim!: continuous_on_eq [OF continuous_on_subset]) | |
| 69620 | 5298 | qed | 
| 5299 | moreover have ?thesis if ?P | |
| 5300 | proof | |
| 5301 | assume ?lhs | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5302 | then obtain c where "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. c) f" | 
| 69620 | 5303 | using homotopic_with_sym by blast | 
| 5304 |     then obtain h where conth: "continuous_on ({0..1::real} \<times> sphere a r) h"
 | |
| 5305 |                     and him: "h ` ({0..1} \<times> sphere a r) \<subseteq> S"
 | |
| 5306 | and h: "\<And>x. h(0, x) = c" "\<And>x. h(1, x) = f x" | |
| 5307 | by (auto simp: homotopic_with_def) | |
| 5308 | obtain b1::'M where "b1 \<in> Basis" | |
| 5309 | using SOME_Basis by auto | |
| 71768 | 5310 |     have "c \<in> h ` ({0..1} \<times> sphere a r)"
 | 
| 5311 | proof | |
| 5312 | show "c = h (0, a + r *\<^sub>R b1)" | |
| 5313 | by (simp add: h) | |
| 5314 |       show "(0, a + r *\<^sub>R b1) \<in> {0..1::real} \<times> sphere a r"
 | |
| 5315 | using greater \<open>b1 \<in> Basis\<close> by (auto simp: dist_norm) | |
| 5316 | qed | |
| 5317 | then have "c \<in> S" | |
| 5318 | using him by blast | |
| 69620 | 5319 |     have uconth: "uniformly_continuous_on ({0..1::real} \<times> (sphere a r)) h"
 | 
| 5320 | by (force intro: compact_Times conth compact_uniformly_continuous) | |
| 5321 | let ?g = "\<lambda>x. h (norm (x - a)/r, | |
| 5322 | a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" | |
| 5323 | let ?g' = "\<lambda>x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" | |
| 5324 | show ?rhs | |
| 5325 | proof (intro exI conjI) | |
| 5326 |       have "continuous_on (cball a r - {a}) ?g'"
 | |
| 72372 | 5327 | using greater | 
| 5328 | by (force simp: dist_norm norm_minus_commute intro: continuous_on_compose2 [OF conth] continuous_intros) | |
| 69620 | 5329 | then show "continuous_on (cball a r) ?g" | 
| 5330 | proof (rule nullhomotopic_from_lemma) | |
| 5331 | show "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> norm (?g' x - ?g a) < e" if "0 < e" for e | |
| 5332 | proof - | |
| 5333 | obtain d where "0 < d" | |
| 72372 | 5334 |              and d: "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> sphere a r; x' \<in> {0..1} \<times> sphere a r; norm ( x' - x) < d\<rbrakk>
 | 
| 5335 | \<Longrightarrow> norm (h x' - h x) < e" | |
| 5336 | using uniformly_continuous_onE [OF uconth \<open>0 < e\<close>] by (auto simp: dist_norm) | |
| 69620 | 5337 | have *: "norm (h (norm (x - a) / r, | 
| 72372 | 5338 | a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" (is "norm (?ha - ?hb) < e") | 
| 69620 | 5339 | if "x \<noteq> a" "norm (x - a) < r" "norm (x - a) < d * r" for x | 
| 5340 | proof - | |
| 72372 | 5341 | have "norm (?ha - ?hb) = norm (?ha - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" | 
| 69620 | 5342 | by (simp add: h) | 
| 5343 | also have "\<dots> < e" | |
| 5344 | using greater \<open>0 < d\<close> \<open>b1 \<in> Basis\<close> that | |
| 72372 | 5345 | by (intro d) (simp_all add: dist_norm, simp add: field_simps) | 
| 69620 | 5346 | finally show ?thesis . | 
| 5347 | qed | |
| 5348 | show ?thesis | |
| 72372 | 5349 | using greater \<open>0 < d\<close> | 
| 5350 | by (rule_tac x = "min r (d * r)" in exI) (auto simp: *) | |
| 69620 | 5351 | qed | 
| 5352 | show "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> ?g x = ?g' x" | |
| 5353 | by auto | |
| 5354 | qed | |
| 5355 | next | |
| 5356 | show "?g ` cball a r \<subseteq> S" | |
| 5357 | using greater him \<open>c \<in> S\<close> | |
| 5358 | by (force simp: h dist_norm norm_minus_commute) | |
| 5359 | next | |
| 5360 | show "\<forall>x\<in>sphere a r. ?g x = f x" | |
| 5361 | using greater by (auto simp: h dist_norm norm_minus_commute) | |
| 5362 | qed | |
| 5363 | next | |
| 5364 | assume ?rhs | |
| 5365 | then obtain g where contg: "continuous_on (cball a r) g" | |
| 5366 | and gim: "g ` cball a r \<subseteq> S" | |
| 5367 | and gf: "\<forall>x \<in> sphere a r. g x = f x" | |
| 5368 | by auto | |
| 5369 | let ?h = "\<lambda>y. g (a + (fst y) *\<^sub>R (snd y - a))" | |
| 5370 |     have "continuous_on ({0..1} \<times> sphere a r) ?h"
 | |
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5371 | proof (rule continuous_on_compose2 [OF contg]) | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5372 |       show "continuous_on ({0..1} \<times> sphere a r) (\<lambda>x. a + fst x *\<^sub>R (snd x - a))"
 | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5373 | by (intro continuous_intros) | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5374 | qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) | 
| 69620 | 5375 | moreover | 
| 5376 |     have "?h ` ({0..1} \<times> sphere a r) \<subseteq> S"
 | |
| 5377 | by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) | |
| 5378 | moreover | |
| 5379 | have "\<forall>x\<in>sphere a r. ?h (0, x) = g a" "\<forall>x\<in>sphere a r. ?h (1, x) = f x" | |
| 5380 | by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) | |
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5381 | ultimately have "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. g a) f" | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5382 | by (auto simp: homotopic_with) | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5383 | then show ?lhs | 
| 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 5384 | using homotopic_with_symD by blast | 
| 69620 | 5385 | qed | 
| 5386 | ultimately | |
| 5387 | show ?thesis by meson | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 5388 | qed | 
| 69620 | 5389 | |
| 5390 | end |