| author | wenzelm | 
| Sat, 08 Sep 2018 22:43:25 +0200 | |
| changeset 68955 | 0851db8cde12 | 
| parent 68607 | 67bb59e49834 | 
| child 69064 | 5840724b1d71 | 
| permissions | -rw-r--r-- | 
| 66835 | 1 | (* Author: L C Paulson, University of Cambridge | 
| 2 | Material split off from Topology_Euclidean_Space | |
| 3 | *) | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 67968 | 5 | section \<open>Connected Components, Homeomorphisms, Baire property, etc\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | theory Connected | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | imports Topology_Euclidean_Space | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | begin | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 67968 | 11 | subsection%unimportant \<open>More properties of closed balls, spheres, etc\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | apply (simp add: interior_def, safe) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | apply (force simp: open_contains_cball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | apply (rule_tac x="ball x e" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | apply (simp add: subset_trans [OF ball_subset_cball]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | lemma islimpt_ball: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 |   fixes x y :: "'a::{real_normed_vector,perfect_space}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | show ?rhs if ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | assume "e \<le> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 |       then have *: "ball x e = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | using ball_eq_empty[of x e] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | have False using \<open>?lhs\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | unfolding * using islimpt_EMPTY[of y] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | then show "e > 0" by (metis not_less) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | show "y \<in> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | ball_subset_cball[of x e] \<open>?lhs\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | unfolding closed_limpt by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | show ?lhs if ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | from that have "e > 0" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | fix d :: real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | assume "d > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | proof (cases "d \<le> dist x y") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | proof (cases "x = y") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | then have False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | using \<open>d \<le> dist x y\<close> \<open>d>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) = | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | unfolding scaleR_minus_left scaleR_one | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | by (auto simp: norm_minus_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | unfolding distrib_right using \<open>x\<noteq>y\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | also have "\<dots> \<le> e - d/2" using \<open>d \<le> dist x y\<close> and \<open>d>0\<close> and \<open>?rhs\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | by (auto simp: dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using \<open>d>0\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | using \<open>x\<noteq>y\<close>[unfolded dist_nz] \<open>d>0\<close> unfolding scaleR_eq_0_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | by (auto simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | unfolding dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | apply simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | unfolding norm_minus_cancel | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | using \<open>d > 0\<close> \<open>x\<noteq>y\<close>[unfolded dist_nz] dist_commute[of x y] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | unfolding dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | apply (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | then have "d > dist x y" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | show "\<exists>x' \<in> ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | proof (cases "x = y") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | obtain z where **: "z \<noteq> y" "dist z y < min e d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | using perfect_choose_dist[of "min e d" y] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | using \<open>d > 0\<close> \<open>e>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | unfolding \<open>x = y\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | using \<open>z \<noteq> y\<close> ** | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | apply (rule_tac x=z in bexI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | apply (auto simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | using \<open>d>0\<close> \<open>d > dist x y\<close> \<open>?rhs\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | apply (rule_tac x=x in bexI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | unfolding mem_cball islimpt_approachable mem_ball by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | lemma closure_ball_lemma: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | fixes x y :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | assumes "x \<noteq> y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | shows "y islimpt ball x (dist x y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | proof (rule islimptI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | fix T | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | assume "y \<in> T" "open T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | unfolding open_dist by fast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | (* choose point between x and y, within distance r of y. *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | define k where "k = min 1 (r / (2 * dist x y))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | define z where "z = y + scaleR k (x - y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | have z_def2: "z = x + scaleR (1 - k) (y - x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | unfolding z_def by (simp add: algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | have "dist z y < r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | unfolding z_def k_def using \<open>0 < r\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | by (simp add: dist_norm min_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | then have "z \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | using \<open>\<forall>z. dist z y < r \<longrightarrow> z \<in> T\<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | have "dist x z < dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | unfolding z_def2 dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | apply (simp add: norm_minus_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | apply (simp only: dist_norm [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | apply (rule mult_strict_right_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | apply (simp add: k_def \<open>0 < r\<close> \<open>x \<noteq> y\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | apply (simp add: \<open>x \<noteq> y\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | then have "z \<in> ball x (dist x y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | have "z \<noteq> y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | unfolding z_def k_def using \<open>x \<noteq> y\<close> \<open>0 < r\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | by (simp add: min_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | using \<open>z \<in> ball x (dist x y)\<close> \<open>z \<in> T\<close> \<open>z \<noteq> y\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | by fast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | |
| 67706 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 158 | lemma at_within_ball_bot_iff: | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 159 |   fixes x y :: "'a::{real_normed_vector,perfect_space}"
 | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 160 | shows "at x within ball y r = bot \<longleftrightarrow> (r=0 \<or> x \<notin> cball y r)" | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 161 | unfolding trivial_limit_within | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 162 | apply (auto simp add:trivial_limit_within islimpt_ball ) | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 163 | by (metis le_less_trans less_eq_real_def zero_le_dist) | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67683diff
changeset | 164 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | lemma closure_ball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | fixes x :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | apply (rule equalityI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | apply (rule closure_minimal) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | apply (rule ball_subset_cball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | apply (rule closed_cball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | apply (rule subsetI, rename_tac y) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | apply (simp add: le_less [where 'a=real]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | apply (erule disjE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | apply (rule subsetD [OF closure_subset], simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | apply (simp add: closure_def, clarify) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | apply (rule closure_ball_lemma) | 
| 66953 | 178 | apply simp | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | (* In a trivial vector space, this fails for e = 0. *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | lemma interior_cball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 |   fixes x :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | shows "interior (cball x e) = ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | proof (cases "e \<ge> 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | case False note cs = this | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 |   from cs have null: "ball x e = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | using ball_empty[of e x] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | assume "y \<in> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | then have False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | by (metis ball_eq_empty null cs dist_eq_0_iff dist_le_zero_iff empty_subsetI mem_cball subset_antisym subset_ball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 |   then have "cball x e = {}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 |   then have "interior (cball x e) = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | using interior_empty by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | ultimately show ?thesis by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | case True note cs = this | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | have "ball x e \<subseteq> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | using ball_subset_cball by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | fix S y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | assume as: "S \<subseteq> cball x e" "open S" "y\<in>S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | then obtain d where "d>0" and d: "\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | unfolding open_dist by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | using perfect_choose_dist [of d] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | have "xa \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | using d[THEN spec[where x = xa]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | using xa by (auto simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | then have xa_cball: "xa \<in> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | using as(1) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | then have "y \<in> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | proof (cases "x = y") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | then have "e > 0" using cs order.order_iff_strict xa_cball xa_y by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | then show "y \<in> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | using \<open>x = y \<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | unfolding dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | using \<open>d>0\<close> norm_ge_zero[of "y - x"] \<open>x \<noteq> y\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | using d as(1)[unfolded subset_eq] by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | have "y - x \<noteq> 0" using \<open>x \<noteq> y\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | hence **:"d / (2 * norm (y - x)) > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | unfolding zero_less_norm_iff[symmetric] using \<open>d>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | by (auto simp: dist_norm algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | by (auto simp: algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | using ** by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | also have "\<dots> = (dist y x) + d/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | using ** by (auto simp: distrib_right dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | finally have "e \<ge> dist x y +d/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | using *[unfolded mem_cball] by (auto simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | then show "y \<in> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | unfolding mem_ball using \<open>d>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | using interior_unique[of "ball x e" "cball x e"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | using open_ball[of x e] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | lemma interior_ball [simp]: "interior (ball x e) = ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | by (simp add: interior_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | lemma frontier_ball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | fixes a :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | shows "0 < e \<Longrightarrow> frontier (ball a e) = sphere a e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by (force simp: frontier_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | lemma frontier_cball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 |   fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | shows "frontier (cball a e) = sphere a e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | by (force simp: frontier_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | lemma cball_eq_empty [simp]: "cball x e = {} \<longleftrightarrow> e < 0"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | apply (simp add: set_eq_iff not_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | apply (metis zero_le_dist dist_self order_less_le_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | lemma cball_empty [simp]: "e < 0 \<Longrightarrow> cball x e = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | lemma cball_eq_sing: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 |   fixes x :: "'a::{metric_space,perfect_space}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 |   shows "cball x e = {x} \<longleftrightarrow> e = 0"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | proof (rule linorder_cases) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | assume e: "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | obtain a where "a \<noteq> x" "dist a x < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | using perfect_choose_dist [OF e] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | then have "a \<noteq> x" "dist x a \<le> e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | by (auto simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | with e show ?thesis by (auto simp: set_eq_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | qed auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | lemma cball_sing: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | fixes x :: "'a::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 |   shows "e = 0 \<Longrightarrow> cball x e = {x}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | by (auto simp: set_eq_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | lemma ball_divide_subset: "d \<ge> 1 \<Longrightarrow> ball x (e/d) \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | apply (cases "e \<le> 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | apply (simp add: ball_empty divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | apply (rule subset_ball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | apply (simp add: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | lemma ball_divide_subset_numeral: "ball x (e / numeral w) \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | using ball_divide_subset one_le_numeral by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | lemma cball_divide_subset: "d \<ge> 1 \<Longrightarrow> cball x (e/d) \<subseteq> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | apply (cases "e < 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | apply (simp add: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | apply (rule subset_cball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | apply (metis div_by_1 frac_le not_le order_refl zero_less_one) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | lemma cball_divide_subset_numeral: "cball x (e / numeral w) \<subseteq> cball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | using cball_divide_subset one_le_numeral by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | lemma compact_cball[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | fixes x :: "'a::heine_borel" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | shows "compact (cball x e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | using compact_eq_bounded_closed bounded_cball closed_cball | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | lemma compact_frontier_bounded[intro]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | fixes S :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | shows "bounded S \<Longrightarrow> compact (frontier S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | unfolding frontier_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | using compact_eq_bounded_closed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | lemma compact_frontier[intro]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | fixes S :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | shows "compact S \<Longrightarrow> compact (frontier S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | using compact_eq_bounded_closed compact_frontier_bounded | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | corollary compact_sphere [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | shows "compact (sphere a r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | using compact_frontier [of "cball a r"] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | corollary bounded_sphere [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | shows "bounded (sphere a r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | by (simp add: compact_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | corollary closed_sphere [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | shows "closed (sphere a r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | by (simp add: compact_imp_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | |
| 67962 | 348 | subsection%unimportant \<open>Connectedness\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | lemma connected_local: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | "connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | \<not> (\<exists>e1 e2. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | openin (subtopology euclidean S) e1 \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | openin (subtopology euclidean S) e2 \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | S \<subseteq> e1 \<union> e2 \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 |       e1 \<inter> e2 = {} \<and>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 |       e1 \<noteq> {} \<and>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 |       e2 \<noteq> {})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | unfolding connected_def openin_open | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | by safe blast+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | lemma exists_diff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | fixes P :: "'a set \<Rightarrow> bool" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | shows "(\<exists>S. P (- S)) \<longleftrightarrow> (\<exists>S. P S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | have ?rhs if ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | using that by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | moreover have "P (- (- S))" if "P S" for S | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | have "S = - (- S)" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | with that show ?thesis by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | ultimately show ?thesis by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | lemma connected_clopen: "connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | (\<forall>T. openin (subtopology euclidean S) T \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 |      closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | have "\<not> connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 |     (\<exists>e1 e2. open e1 \<and> open (- e2) \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | unfolding connected_def openin_open closedin_closed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | by (metis double_complement) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | then have th0: "connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 |     \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | by (simp add: closed_def) metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 |   have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | unfolding connected_def openin_open closedin_closed by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" for e2 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 |     have "?P e2 e1 \<longleftrightarrow> (\<exists>t. closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t \<noteq> S)" for e1
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | by (simp add: th0 th1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | lemma connected_linear_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | assumes "linear f" and "connected s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | shows "connected (f ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | subsection \<open>Connected components, considered as a connectedness relation or a set\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | |
| 67962 | 413 | definition%important "connected_component s x y \<equiv> \<exists>t. connected t \<and> t \<subseteq> s \<and> x \<in> t \<and> y \<in> t" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | abbreviation "connected_component_set s x \<equiv> Collect (connected_component s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | lemma connected_componentI: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | "connected t \<Longrightarrow> t \<subseteq> s \<Longrightarrow> x \<in> t \<Longrightarrow> y \<in> t \<Longrightarrow> connected_component s x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | by (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | lemma connected_component_in: "connected_component s x y \<Longrightarrow> x \<in> s \<and> y \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | by (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | lemma connected_component_refl: "x \<in> s \<Longrightarrow> connected_component s x x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | by (auto simp: connected_component_def) (use connected_sing in blast) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | lemma connected_component_refl_eq [simp]: "connected_component s x x \<longleftrightarrow> x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | by (auto simp: connected_component_refl) (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | lemma connected_component_sym: "connected_component s x y \<Longrightarrow> connected_component s y x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | by (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | lemma connected_component_trans: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | "connected_component s x y \<Longrightarrow> connected_component s y z \<Longrightarrow> connected_component s x z" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | unfolding connected_component_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | by (metis Int_iff Un_iff Un_subset_iff equals0D connected_Un) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | lemma connected_component_of_subset: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | "connected_component s x y \<Longrightarrow> s \<subseteq> t \<Longrightarrow> connected_component t x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | by (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | lemma connected_component_Union: "connected_component_set s x = \<Union>{t. connected t \<and> x \<in> t \<and> t \<subseteq> s}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | by (auto simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | lemma connected_connected_component [iff]: "connected (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | by (auto simp: connected_component_Union intro: connected_Union) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | lemma connected_iff_eq_connected_component_set: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 | "connected s \<longleftrightarrow> (\<forall>x \<in> s. connected_component_set s x = s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | proof (cases "s = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 | then show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | then obtain x where "x \<in> s" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | assume "connected s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | then show "\<forall>x \<in> s. connected_component_set s x = s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | by (force simp: connected_component_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | assume "\<forall>x \<in> s. connected_component_set s x = s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | then show "connected s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | by (metis \<open>x \<in> s\<close> connected_connected_component) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | lemma connected_component_subset: "connected_component_set s x \<subseteq> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | using connected_component_in by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | lemma connected_component_eq_self: "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> connected_component_set s x = s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | by (simp add: connected_iff_eq_connected_component_set) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | lemma connected_iff_connected_component: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | "connected s \<longleftrightarrow> (\<forall>x \<in> s. \<forall>y \<in> s. connected_component s x y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | using connected_component_in by (auto simp: connected_iff_eq_connected_component_set) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | lemma connected_component_maximal: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | "x \<in> t \<Longrightarrow> connected t \<Longrightarrow> t \<subseteq> s \<Longrightarrow> t \<subseteq> (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | using connected_component_eq_self connected_component_of_subset by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | lemma connected_component_mono: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | "s \<subseteq> t \<Longrightarrow> connected_component_set s x \<subseteq> connected_component_set t x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | by (simp add: Collect_mono connected_component_of_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 | lemma connected_component_eq_empty [simp]: "connected_component_set s x = {} \<longleftrightarrow> x \<notin> s"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | using connected_component_refl by (fastforce simp: connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | lemma connected_component_set_empty [simp]: "connected_component_set {} x = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | using connected_component_eq_empty by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | lemma connected_component_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | "y \<in> connected_component_set s x \<Longrightarrow> (connected_component_set s y = connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | by (metis (no_types, lifting) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | Collect_cong connected_component_sym connected_component_trans mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | lemma closed_connected_component: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | assumes s: "closed s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | shows "closed (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | proof (cases "x \<in> s") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | by (metis connected_component_eq_empty closed_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 506 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 507 | unfolding closure_eq [symmetric] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | show "closure (connected_component_set s x) \<subseteq> connected_component_set s x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | apply (rule connected_component_maximal) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | apply (simp add: closure_def True) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | apply (simp add: connected_imp_connected_closure) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | apply (simp add: s closure_minimal connected_component_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 514 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 515 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | show "connected_component_set s x \<subseteq> closure (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 517 | by (simp add: closure_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | lemma connected_component_disjoint: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 |   "connected_component_set s a \<inter> connected_component_set s b = {} \<longleftrightarrow>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | a \<notin> connected_component_set s b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | apply (auto simp: connected_component_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 525 | using connected_component_eq connected_component_sym | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | lemma connected_component_nonoverlap: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 |   "connected_component_set s a \<inter> connected_component_set s b = {} \<longleftrightarrow>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 531 | a \<notin> s \<or> b \<notin> s \<or> connected_component_set s a \<noteq> connected_component_set s b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 | apply (auto simp: connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 533 | using connected_component_refl_eq | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | apply (metis connected_component_eq mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | apply (metis connected_component_eq mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 538 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | lemma connected_component_overlap: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 |   "connected_component_set s a \<inter> connected_component_set s b \<noteq> {} \<longleftrightarrow>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 541 | a \<in> s \<and> b \<in> s \<and> connected_component_set s a = connected_component_set s b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | by (auto simp: connected_component_nonoverlap) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 | lemma connected_component_sym_eq: "connected_component s x y \<longleftrightarrow> connected_component s y x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 545 | using connected_component_sym by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 | lemma connected_component_eq_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 | "connected_component_set s x = connected_component_set s y \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 549 | x \<notin> s \<and> y \<notin> s \<or> x \<in> s \<and> y \<in> s \<and> connected_component s x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | apply (cases "y \<in> s", simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | apply (metis connected_component_eq connected_component_eq_empty connected_component_refl_eq mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 | apply (cases "x \<in> s", simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | apply (metis connected_component_eq_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 | using connected_component_eq_empty | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 | apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | lemma connected_iff_connected_component_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | "connected s \<longleftrightarrow> (\<forall>x \<in> s. \<forall>y \<in> s. connected_component_set s x = connected_component_set s y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | by (simp add: connected_component_eq_eq connected_iff_connected_component) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 | lemma connected_component_idemp: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | "connected_component_set (connected_component_set s x) x = connected_component_set s x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | apply (rule subset_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | apply (simp add: connected_component_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | apply (metis connected_component_eq_empty connected_component_maximal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 | connected_component_refl_eq connected_connected_component mem_Collect_eq set_eq_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | lemma connected_component_unique: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | "\<lbrakk>x \<in> c; c \<subseteq> s; connected c; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | \<And>c'. x \<in> c' \<and> c' \<subseteq> s \<and> connected c' | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | \<Longrightarrow> c' \<subseteq> c\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | \<Longrightarrow> connected_component_set s x = c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 | apply (rule subset_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | apply (meson connected_component_maximal connected_component_subset connected_connected_component contra_subsetD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 | by (simp add: connected_component_maximal) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | lemma joinable_connected_component_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | "\<lbrakk>connected t; t \<subseteq> s; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 |     connected_component_set s x \<inter> t \<noteq> {};
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 |     connected_component_set s y \<inter> t \<noteq> {}\<rbrakk>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | \<Longrightarrow> connected_component_set s x = connected_component_set s y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | apply (simp add: ex_in_conv [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | apply (rule connected_component_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | by (metis (no_types, hide_lams) connected_component_eq_eq connected_component_in connected_component_maximal subsetD mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | lemma Union_connected_component: "\<Union>(connected_component_set s ` s) = s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | apply (rule subset_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | apply (simp add: SUP_least connected_component_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | using connected_component_refl_eq | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | lemma complement_connected_component_unions: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | "s - connected_component_set s x = | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 |      \<Union>(connected_component_set s ` s - {connected_component_set s x})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 | apply (subst Union_connected_component [symmetric], auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | apply (metis connected_component_eq_eq connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | by (metis connected_component_eq mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | lemma connected_component_intermediate_subset: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | "\<lbrakk>connected_component_set u a \<subseteq> t; t \<subseteq> u\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | \<Longrightarrow> connected_component_set t a = connected_component_set u a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | apply (case_tac "a \<in> u") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | apply (simp add: connected_component_maximal connected_component_mono subset_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | using connected_component_eq_empty by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | subsection \<open>The set of connected components of a set\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | |
| 67962 | 613 | definition%important components:: "'a::topological_space set \<Rightarrow> 'a set set" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | where "components s \<equiv> connected_component_set s ` s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 | lemma components_iff: "s \<in> components u \<longleftrightarrow> (\<exists>x. x \<in> u \<and> s = connected_component_set u x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | by (auto simp: components_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | lemma componentsI: "x \<in> u \<Longrightarrow> connected_component_set u x \<in> components u" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 | by (auto simp: components_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 | lemma componentsE: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | assumes "s \<in> components u" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | obtains x where "x \<in> u" "s = connected_component_set u x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | using assms by (auto simp: components_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | lemma Union_components [simp]: "\<Union>(components u) = u" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | apply (rule subset_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | using Union_connected_component components_def apply fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | apply (metis Union_connected_component components_def set_eq_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | lemma pairwise_disjoint_components: "pairwise (\<lambda>X Y. X \<inter> Y = {}) (components u)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | apply (simp add: pairwise_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | apply (auto simp: components_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | apply (metis connected_component_eq_eq connected_component_in)+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | lemma in_components_nonempty: "c \<in> components s \<Longrightarrow> c \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | by (metis components_iff connected_component_eq_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | lemma in_components_subset: "c \<in> components s \<Longrightarrow> c \<subseteq> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | using Union_components by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | lemma in_components_connected: "c \<in> components s \<Longrightarrow> connected c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | by (metis components_iff connected_connected_component) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | lemma in_components_maximal: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | "c \<in> components s \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 |     c \<noteq> {} \<and> c \<subseteq> s \<and> connected c \<and> (\<forall>d. d \<noteq> {} \<and> c \<subseteq> d \<and> d \<subseteq> s \<and> connected d \<longrightarrow> d = c)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | apply (rule iffI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | apply (simp add: in_components_nonempty in_components_connected) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | apply (metis (full_types) components_iff connected_component_eq_self connected_component_intermediate_subset connected_component_refl in_components_subset mem_Collect_eq rev_subsetD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | apply (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | lemma joinable_components_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 |   "connected t \<and> t \<subseteq> s \<and> c1 \<in> components s \<and> c2 \<in> components s \<and> c1 \<inter> t \<noteq> {} \<and> c2 \<inter> t \<noteq> {} \<Longrightarrow> c1 = c2"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | by (metis (full_types) components_iff joinable_connected_component_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 | lemma closed_components: "\<lbrakk>closed s; c \<in> components s\<rbrakk> \<Longrightarrow> closed c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | by (metis closed_connected_component components_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | lemma compact_components: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 | fixes s :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 | shows "\<lbrakk>compact s; c \<in> components s\<rbrakk> \<Longrightarrow> compact c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | by (meson bounded_subset closed_components in_components_subset compact_eq_bounded_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 669 | lemma components_nonoverlap: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 |     "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c \<inter> c' = {}) \<longleftrightarrow> (c \<noteq> c')"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | apply (auto simp: in_components_nonempty components_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | using connected_component_refl apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | apply (metis connected_component_eq_eq connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | by (metis connected_component_eq mem_Collect_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | lemma components_eq: "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c = c' \<longleftrightarrow> c \<inter> c' \<noteq> {})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | by (metis components_nonoverlap) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | lemma components_eq_empty [simp]: "components s = {} \<longleftrightarrow> s = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | by (simp add: components_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | lemma components_empty [simp]: "components {} = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | lemma connected_eq_connected_components_eq: "connected s \<longleftrightarrow> (\<forall>c \<in> components s. \<forall>c' \<in> components s. c = c')" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | by (metis (no_types, hide_lams) components_iff connected_component_eq_eq connected_iff_connected_component) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | lemma components_eq_sing_iff: "components s = {s} \<longleftrightarrow> connected s \<and> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | apply (rule iffI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | using in_components_connected apply fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | apply safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 | using Union_components apply fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | apply (metis components_iff connected_component_eq_self) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | using in_components_maximal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | lemma components_eq_sing_exists: "(\<exists>a. components s = {a}) \<longleftrightarrow> connected s \<and> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | apply (rule iffI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | using connected_eq_connected_components_eq apply fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | apply (metis components_eq_sing_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | lemma connected_eq_components_subset_sing: "connected s \<longleftrightarrow> components s \<subseteq> {s}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 | by (metis Union_components components_empty components_eq_sing_iff connected_empty insert_subset order_refl subset_singletonD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | lemma connected_eq_components_subset_sing_exists: "connected s \<longleftrightarrow> (\<exists>a. components s \<subseteq> {a})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | by (metis components_eq_sing_exists connected_eq_components_subset_sing empty_iff subset_iff subset_singletonD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | lemma in_components_self: "s \<in> components s \<longleftrightarrow> connected s \<and> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | by (metis components_empty components_eq_sing_iff empty_iff in_components_connected insertI1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | lemma components_maximal: "\<lbrakk>c \<in> components s; connected t; t \<subseteq> s; c \<inter> t \<noteq> {}\<rbrakk> \<Longrightarrow> t \<subseteq> c"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 | apply (simp add: components_def ex_in_conv [symmetric], clarify) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | by (meson connected_component_def connected_component_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | lemma exists_component_superset: "\<lbrakk>t \<subseteq> s; s \<noteq> {}; connected t\<rbrakk> \<Longrightarrow> \<exists>c. c \<in> components s \<and> t \<subseteq> c"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 |   apply (cases "t = {}", force)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | apply (metis components_def ex_in_conv connected_component_maximal contra_subsetD image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | lemma components_intermediate_subset: "\<lbrakk>s \<in> components u; s \<subseteq> t; t \<subseteq> u\<rbrakk> \<Longrightarrow> s \<in> components t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | apply (auto simp: components_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | apply (metis connected_component_eq_empty connected_component_intermediate_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | lemma in_components_unions_complement: "c \<in> components s \<Longrightarrow> s - c = \<Union>(components s - {c})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 | by (metis complement_connected_component_unions components_def components_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | lemma connected_intermediate_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | assumes cs: "connected s" and st: "s \<subseteq> t" and ts: "t \<subseteq> closure s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | shows "connected t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | proof (rule connectedI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | fix A B | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 |   assume A: "open A" and B: "open B" and Alap: "A \<inter> t \<noteq> {}" and Blap: "B \<inter> t \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 |     and disj: "A \<inter> B \<inter> t = {}" and cover: "t \<subseteq> A \<union> B"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 |   have disjs: "A \<inter> B \<inter> s = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | using disj st by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 |   have "A \<inter> closure s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | using Alap Int_absorb1 ts by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 |   then have Alaps: "A \<inter> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | by (simp add: A open_Int_closure_eq_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 |   have "B \<inter> closure s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | using Blap Int_absorb1 ts by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 |   then have Blaps: "B \<inter> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | by (simp add: B open_Int_closure_eq_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | then show False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | using cs [unfolded connected_def] A B disjs Alaps Blaps cover st | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | lemma closedin_connected_component: "closedin (subtopology euclidean s) (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 753 | proof (cases "connected_component_set s x = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 | by (metis closedin_empty) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 | then obtain y where y: "connected_component s x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 760 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 | have *: "connected_component_set s x \<subseteq> s \<inter> closure (connected_component_set s x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | by (auto simp: closure_def connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 | have "connected_component s x y \<Longrightarrow> s \<inter> closure (connected_component_set s x) \<subseteq> connected_component_set s x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 764 | apply (rule connected_component_maximal, simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 765 | using closure_subset connected_component_in apply fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 | using * connected_intermediate_closure apply blast+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 767 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 768 | with y * show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 769 | by (auto simp: closedin_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 770 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 771 | |
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 772 | lemma closedin_component: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 773 | "C \<in> components s \<Longrightarrow> closedin (subtopology euclidean s) C" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 774 | using closedin_connected_component componentsE by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 775 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 776 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | subsection \<open>Intersecting chains of compact sets and the Baire property\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 779 | proposition bounded_closed_chain: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | fixes \<F> :: "'a::heine_borel set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 |   assumes "B \<in> \<F>" "bounded B" and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" and "{} \<notin> \<F>"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | and chain: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 |     shows "\<Inter>\<F> \<noteq> {}"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 784 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 785 |   have "B \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 | proof (rule compact_imp_fip) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | show "compact B" "\<And>T. T \<in> \<F> \<Longrightarrow> closed T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 788 | by (simp_all add: assms compact_eq_bounded_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 |     show "\<lbrakk>finite \<G>; \<G> \<subseteq> \<F>\<rbrakk> \<Longrightarrow> B \<inter> \<Inter>\<G> \<noteq> {}" for \<G>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 790 | proof (induction \<G> rule: finite_induct) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | case empty | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 792 | with assms show ?case by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 793 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | case (insert U \<G>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 795 |       then have "U \<in> \<F>" and ne: "B \<inter> \<Inter>\<G> \<noteq> {}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | then consider "B \<subseteq> U" | "U \<subseteq> B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | using \<open>B \<in> \<F>\<close> chain by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 | then show ?case | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | proof cases | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | case 1 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | using Int_left_commute ne by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 803 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 | case 2 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 805 |           have "U \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 |             using \<open>U \<in> \<F>\<close> \<open>{} \<notin> \<F>\<close> by blast
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 807 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 808 | have False if "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. x \<notin> Y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 809 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 | have "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. Y \<subseteq> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 | by (metis chain contra_subsetD insert.prems insert_subset that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 | then obtain Y where "Y \<in> \<G>" "Y \<subseteq> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 |               by (metis all_not_in_conv \<open>U \<noteq> {}\<close>)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | moreover obtain x where "x \<in> \<Inter>\<G>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | by (metis Int_emptyI ne) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 | by (metis Inf_lower subset_eq that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 818 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 819 | with 2 show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 821 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 822 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 823 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | then show ?thesis by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 827 | corollary compact_chain: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | fixes \<F> :: "'a::heine_borel set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 |   assumes "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" "{} \<notin> \<F>"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 830 | "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 831 |     shows "\<Inter> \<F> \<noteq> {}"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 832 | proof (cases "\<F> = {}")
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 833 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 834 | then show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 838 | by (metis False all_not_in_conv assms compact_imp_bounded compact_imp_closed bounded_closed_chain) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 839 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 840 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 841 | lemma compact_nest: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 | fixes F :: "'a::linorder \<Rightarrow> 'b::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 843 |   assumes F: "\<And>n. compact(F n)" "\<And>n. F n \<noteq> {}" and mono: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 844 |   shows "\<Inter>range F \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 846 | have *: "\<And>S T. S \<in> range F \<and> T \<in> range F \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | by (metis mono image_iff le_cases) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 | apply (rule compact_chain [OF _ _ *]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 | using F apply (blast intro: dest: *)+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 851 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 852 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 853 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 | text\<open>The Baire property of dense sets\<close> | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 855 | theorem Baire: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 856 |   fixes S::"'a::{real_normed_vector,heine_borel} set"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 857 | assumes "closed S" "countable \<G>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 | and ope: "\<And>T. T \<in> \<G> \<Longrightarrow> openin (subtopology euclidean S) T \<and> S \<subseteq> closure T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 | shows "S \<subseteq> closure(\<Inter>\<G>)" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 860 | proof (cases "\<G> = {}")
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | using closure_subset by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | let ?g = "from_nat_into \<G>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | then have gin: "?g n \<in> \<G>" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 868 | by (simp add: from_nat_into) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 869 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 870 | proof (clarsimp simp: closure_approachable) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 | fix x and e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 872 | assume "x \<in> S" "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 873 | obtain TF where opeF: "\<And>n. openin (subtopology euclidean S) (TF n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 |                and ne: "\<And>n. TF n \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 875 | and subg: "\<And>n. S \<inter> closure(TF n) \<subseteq> ?g n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | and subball: "\<And>n. closure(TF n) \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | and decr: "\<And>n. TF(Suc n) \<subseteq> TF n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 878 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 |       have *: "\<exists>Y. (openin (subtopology euclidean S) Y \<and> Y \<noteq> {} \<and>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 880 | S \<inter> closure Y \<subseteq> ?g n \<and> closure Y \<subseteq> ball x e) \<and> Y \<subseteq> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 881 |         if opeU: "openin (subtopology euclidean S) U" and "U \<noteq> {}" and cloU: "closure U \<subseteq> ball x e" for U n
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 882 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | obtain T where T: "open T" "U = T \<inter> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | using \<open>openin (subtopology euclidean S) U\<close> by (auto simp: openin_subtopology) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 |         with \<open>U \<noteq> {}\<close> have "T \<inter> closure (?g n) \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 886 | using gin ope by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 |         then have "T \<inter> ?g n \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 888 | using \<open>open T\<close> open_Int_closure_eq_empty by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | then obtain y where "y \<in> U" "y \<in> ?g n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 890 | using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 891 | moreover have "openin (subtopology euclidean S) (U \<inter> ?g n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 892 | using gin ope opeU by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 893 | ultimately obtain d where U: "U \<inter> ?g n \<subseteq> S" and "d > 0" and d: "ball y d \<inter> S \<subseteq> U \<inter> ?g n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | by (force simp: openin_contains_ball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 895 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 896 | proof (intro exI conjI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 897 | show "openin (subtopology euclidean S) (S \<inter> ball y (d/2))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 | by (simp add: openin_open_Int) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 899 |           show "S \<inter> ball y (d/2) \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | using \<open>0 < d\<close> \<open>y \<in> U\<close> opeU openin_imp_subset by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 901 | have "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> closure (ball y (d/2))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 902 | using closure_mono by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | also have "... \<subseteq> ?g n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 | using \<open>d > 0\<close> d by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 905 | finally show "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> ?g n" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | have "closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> ball y d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 907 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | have "closure (ball y (d/2)) \<subseteq> ball y d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 909 | using \<open>d > 0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 910 | then have "closure (S \<inter> ball y (d/2)) \<subseteq> ball y d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 911 | by (meson closure_mono inf.cobounded2 subset_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 913 | by (simp add: \<open>closed S\<close> closure_minimal) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 915 | also have "... \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | using cloU closure_subset d by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 | finally show "closure (S \<inter> ball y (d/2)) \<subseteq> ball x e" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 918 | show "S \<inter> ball y (d/2) \<subseteq> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 919 | using ball_divide_subset_numeral d by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 921 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 922 |       let ?\<Phi> = "\<lambda>n X. openin (subtopology euclidean S) X \<and> X \<noteq> {} \<and>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 923 | S \<inter> closure X \<subseteq> ?g n \<and> closure X \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 924 | have "closure (S \<inter> ball x (e / 2)) \<subseteq> closure(ball x (e/2))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 925 | by (simp add: closure_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | also have "... \<subseteq> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | using \<open>e > 0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 928 | finally have "closure (S \<inter> ball x (e / 2)) \<subseteq> ball x e" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 929 |       moreover have"openin (subtopology euclidean S) (S \<inter> ball x (e / 2))" "S \<inter> ball x (e / 2) \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | using \<open>0 < e\<close> \<open>x \<in> S\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | ultimately obtain Y where Y: "?\<Phi> 0 Y \<and> Y \<subseteq> S \<inter> ball x (e / 2)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 932 | using * [of "S \<inter> ball x (e/2)" 0] by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | show thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 934 | proof (rule exE [OF dependent_nat_choice [of ?\<Phi> "\<lambda>n X Y. Y \<subseteq> X"]]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 935 | show "\<exists>x. ?\<Phi> 0 x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 936 | using Y by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 937 | show "\<exists>Y. ?\<Phi> (Suc n) Y \<and> Y \<subseteq> X" if "?\<Phi> n X" for X n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 | using that by (blast intro: *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | qed (use that in metis) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 941 |     have "(\<Inter>n. S \<inter> closure (TF n)) \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | proof (rule compact_nest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 943 | show "\<And>n. compact (S \<inter> closure (TF n))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 944 | by (metis closed_closure subball bounded_subset_ballI compact_eq_bounded_closed closed_Int_compact [OF \<open>closed S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 945 |       show "\<And>n. S \<inter> closure (TF n) \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 | by (metis Int_absorb1 opeF \<open>closed S\<close> closure_eq_empty closure_minimal ne openin_imp_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 947 | show "\<And>m n. m \<le> n \<Longrightarrow> S \<inter> closure (TF n) \<subseteq> S \<inter> closure (TF m)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 | by (meson closure_mono decr dual_order.refl inf_mono lift_Suc_antimono_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 949 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 |     moreover have "(\<Inter>n. S \<inter> closure (TF n)) \<subseteq> {y \<in> \<Inter>\<G>. dist y x < e}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 951 | proof (clarsimp, intro conjI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 952 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 953 | assume "y \<in> S" and y: "\<forall>n. y \<in> closure (TF n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 954 | then show "\<forall>T\<in>\<G>. y \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 955 | by (metis Int_iff from_nat_into_surj [OF \<open>countable \<G>\<close>] set_mp subg) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 956 | show "dist y x < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 957 | by (metis y dist_commute mem_ball subball subsetCE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 959 | ultimately show "\<exists>y \<in> \<Inter>\<G>. dist y x < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 960 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 961 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 962 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 963 | |
| 67968 | 964 | subsection%unimportant \<open>Some theorems on sups and infs using the notion "bounded"\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 965 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 966 | lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 967 | by (simp add: bounded_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 968 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 969 | lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | by (auto simp: bounded_def bdd_above_def dist_real_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 | (metis abs_le_D1 abs_minus_commute diff_le_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 972 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 974 | by (auto simp: bounded_def bdd_below_def dist_real_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 975 | (metis abs_le_D1 add.commute diff_le_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 976 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | lemma bounded_inner_imp_bdd_above: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | assumes "bounded s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 979 | shows "bdd_above ((\<lambda>x. x \<bullet> a) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 980 | by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 982 | lemma bounded_inner_imp_bdd_below: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | assumes "bounded s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 984 | shows "bdd_below ((\<lambda>x. x \<bullet> a) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 985 | by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 986 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 987 | lemma bounded_has_Sup: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | fixes S :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 989 | assumes "bounded S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 990 |     and "S \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 991 | shows "\<forall>x\<in>S. x \<le> Sup S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 992 | and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 993 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 994 | show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 995 | using assms by (metis cSup_least) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 996 | qed (metis cSup_upper assms(1) bounded_imp_bdd_above) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 997 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 998 | lemma Sup_insert: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | fixes S :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 |   shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1001 | by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1002 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | lemma Sup_insert_finite: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1004 | fixes S :: "'a::conditionally_complete_linorder set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1005 |   shows "finite S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 | by (simp add: cSup_insert sup_max) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1008 | lemma bounded_has_Inf: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1009 | fixes S :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 | assumes "bounded S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 |     and "S \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | shows "\<forall>x\<in>S. x \<ge> Inf S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1013 | and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1015 | show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1016 | using assms by (metis cInf_greatest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1017 | qed (metis cInf_lower assms(1) bounded_imp_bdd_below) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1018 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1019 | lemma Inf_insert: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1020 | fixes S :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1021 |   shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1022 | by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1024 | lemma Inf_insert_finite: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1025 | fixes S :: "'a::conditionally_complete_linorder set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1026 |   shows "finite S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1027 | by (simp add: cInf_eq_Min) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1029 | lemma finite_imp_less_Inf: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | fixes a :: "'a::conditionally_complete_linorder" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1031 | shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a < x\<rbrakk> \<Longrightarrow> a < Inf X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 | lemma finite_less_Inf_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 | fixes a :: "'a :: conditionally_complete_linorder" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1036 |   shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a < Inf X \<longleftrightarrow> (\<forall>x \<in> X. a < x)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1037 | by (auto simp: cInf_eq_Min) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1038 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1039 | lemma finite_imp_Sup_less: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1040 | fixes a :: "'a::conditionally_complete_linorder" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1041 | shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a > x\<rbrakk> \<Longrightarrow> a > Sup X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1042 | by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1044 | lemma finite_Sup_less_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1045 | fixes a :: "'a :: conditionally_complete_linorder" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1046 |   shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a > Sup X \<longleftrightarrow> (\<forall>x \<in> X. a > x)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1047 | by (auto simp: cSup_eq_Max) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1048 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1049 | proposition is_interval_compact: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1050 | "is_interval S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = cbox a b)" (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | proof (cases "S = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 | with empty_as_interval show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1054 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1055 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1058 | assume L: ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | then have "is_interval S" "compact S" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1060 | define a where "a \<equiv> \<Sum>i\<in>Basis. (INF x:S. x \<bullet> i) *\<^sub>R i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | define b where "b \<equiv> \<Sum>i\<in>Basis. (SUP x:S. x \<bullet> i) *\<^sub>R i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1062 | have 1: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> (INF x:S. x \<bullet> i) \<le> x \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1063 | by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | have 2: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> x \<bullet> i \<le> (SUP x:S. x \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1065 | by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1066 | have 3: "x \<in> S" if inf: "\<And>i. i \<in> Basis \<Longrightarrow> (INF x:S. x \<bullet> i) \<le> x \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1067 | and sup: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<le> (SUP x:S. x \<bullet> i)" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | proof (rule mem_box_componentwiseI [OF \<open>is_interval S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | fix i::'a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | assume i: "i \<in> Basis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1071 | have cont: "continuous_on S (\<lambda>x. x \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | by (intro continuous_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | obtain a where "a \<in> S" and a: "\<And>y. y\<in>S \<Longrightarrow> a \<bullet> i \<le> y \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 | using continuous_attains_inf [OF \<open>compact S\<close> False cont] by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1075 | obtain b where "b \<in> S" and b: "\<And>y. y\<in>S \<Longrightarrow> y \<bullet> i \<le> b \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1076 | using continuous_attains_sup [OF \<open>compact S\<close> False cont] by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1077 | have "a \<bullet> i \<le> (INF x:S. x \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1078 | by (simp add: False a cINF_greatest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | also have "\<dots> \<le> x \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | by (simp add: i inf) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 | finally have ai: "a \<bullet> i \<le> x \<bullet> i" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1082 | have "x \<bullet> i \<le> (SUP x:S. x \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1083 | by (simp add: i sup) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | also have "(SUP x:S. x \<bullet> i) \<le> b \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | by (simp add: False b cSUP_least) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1086 | finally have bi: "x \<bullet> i \<le> b \<bullet> i" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1087 | show "x \<bullet> i \<in> (\<lambda>x. x \<bullet> i) ` S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1088 | apply (rule_tac x="\<Sum>j\<in>Basis. (if j = i then x \<bullet> i else a \<bullet> j) *\<^sub>R j" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1089 | apply (simp add: i) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1090 | apply (rule mem_is_intervalI [OF \<open>is_interval S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1091 | using i ai bi apply force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1092 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1093 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1094 | have "S = cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1095 | by (auto simp: a_def b_def mem_box intro: 1 2 3) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1097 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1099 | assume R: ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1100 | then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 | using compact_cbox is_interval_cbox by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1104 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1105 | text \<open>Hence some handy theorems on distance, diameter etc. of/from a set.\<close> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1106 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1107 | lemma distance_attains_sup: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1108 |   assumes "compact s" "s \<noteq> {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1109 | shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a y \<le> dist a x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1110 | proof (rule continuous_attains_sup [OF assms]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1111 |   {
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1112 | fix x | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1113 | assume "x\<in>s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1114 | have "(dist a \<longlongrightarrow> dist a x) (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1115 | by (intro tendsto_dist tendsto_const tendsto_ident_at) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1116 | } | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1117 | then show "continuous_on s (dist a)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1118 | unfolding continuous_on .. | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1119 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1120 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1121 | text \<open>For \emph{minimal} distance, we only need closure, not compactness.\<close>
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1122 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1123 | lemma distance_attains_inf: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1124 | fixes a :: "'a::heine_borel" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1125 |   assumes "closed s" and "s \<noteq> {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1126 | obtains x where "x\<in>s" "\<And>y. y \<in> s \<Longrightarrow> dist a x \<le> dist a y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1127 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1128 | from assms obtain b where "b \<in> s" by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1129 | let ?B = "s \<inter> cball a (dist b a)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1130 |   have "?B \<noteq> {}" using \<open>b \<in> s\<close>
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1131 | by (auto simp: dist_commute) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1132 | moreover have "continuous_on ?B (dist a)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1133 | by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_ident continuous_const) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1134 | moreover have "compact ?B" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1135 | by (intro closed_Int_compact \<open>closed s\<close> compact_cball) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1136 | ultimately obtain x where "x \<in> ?B" "\<forall>y\<in>?B. dist a x \<le> dist a y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1137 | by (metis continuous_attains_inf) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1138 | with that show ?thesis by fastforce | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1139 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1140 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1141 | |
| 67968 | 1142 | subsection%unimportant\<open>Relations among convergence and absolute convergence for power series\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1144 | lemma summable_imp_bounded: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1146 | shows "summable f \<Longrightarrow> bounded (range f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | by (frule summable_LIMSEQ_zero) (simp add: convergent_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1148 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | lemma summable_imp_sums_bounded: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1150 |    "summable f \<Longrightarrow> bounded (range (\<lambda>n. sum f {..<n}))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1151 | by (auto simp: summable_def sums_def dest: convergent_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1152 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1153 | lemma power_series_conv_imp_absconv_weak: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1154 |   fixes a:: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}" and w :: 'a
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1155 | assumes sum: "summable (\<lambda>n. a n * z ^ n)" and no: "norm w < norm z" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1156 | shows "summable (\<lambda>n. of_real(norm(a n)) * w ^ n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1157 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1158 | obtain M where M: "\<And>x. norm (a x * z ^ x) \<le> M" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1159 | using summable_imp_bounded [OF sum] by (force simp: bounded_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1160 | then have *: "summable (\<lambda>n. norm (a n) * norm w ^ n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1161 | by (rule_tac M=M in Abel_lemma) (auto simp: norm_mult norm_power intro: no) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1162 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 | apply (rule series_comparison_complex [of "(\<lambda>n. of_real(norm(a n) * norm w ^ n))"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 | apply (simp only: summable_complex_of_real *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | apply (auto simp: norm_mult norm_power) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1166 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1167 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | |
| 67962 | 1169 | subsection%unimportant \<open>Bounded closed nest property (proof does not use Heine-Borel)\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 | lemma bounded_closed_nest: | 
| 68302 | 1172 |   fixes S :: "nat \<Rightarrow> ('a::heine_borel) set"
 | 
| 1173 | assumes "\<And>n. closed (S n)" | |
| 1174 |       and "\<And>n. S n \<noteq> {}"
 | |
| 1175 | and "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m" | |
| 1176 | and "bounded (S 0)" | |
| 1177 | obtains a where "\<And>n. a \<in> S n" | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | proof - | 
| 68302 | 1179 | from assms(2) obtain x where x: "\<forall>n. x n \<in> S n" | 
| 1180 | using choice[of "\<lambda>n x. x \<in> S n"] by auto | |
| 1181 | from assms(4,1) have "seq_compact (S 0)" | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 | by (simp add: bounded_closed_imp_seq_compact) | 
| 68302 | 1183 | then obtain l r where lr: "l \<in> S 0" "strict_mono r" "(x \<circ> r) \<longlonglongrightarrow> l" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1184 | using x and assms(3) unfolding seq_compact_def by blast | 
| 68302 | 1185 | have "\<forall>n. l \<in> S n" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1186 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1187 | fix n :: nat | 
| 68302 | 1188 | have "closed (S n)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | using assms(1) by simp | 
| 68302 | 1190 | moreover have "\<forall>i. (x \<circ> r) i \<in> S i" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1191 | using x and assms(3) and lr(2) [THEN seq_suble] by auto | 
| 68302 | 1192 | then have "\<forall>i. (x \<circ> r) (i + n) \<in> S n" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | using assms(3) by (fast intro!: le_add2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | moreover have "(\<lambda>i. (x \<circ> r) (i + n)) \<longlonglongrightarrow> l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | using lr(3) by (rule LIMSEQ_ignore_initial_segment) | 
| 68302 | 1196 | ultimately show "l \<in> S n" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | by (rule closed_sequentially) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1198 | qed | 
| 68302 | 1199 | then show ?thesis | 
| 1200 | using that by blast | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1203 | text \<open>Decreasing case does not even need compactness, just completeness.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | lemma decreasing_closed_nest: | 
| 68302 | 1206 |   fixes S :: "nat \<Rightarrow> ('a::complete_space) set"
 | 
| 1207 | assumes "\<And>n. closed (S n)" | |
| 1208 |           "\<And>n. S n \<noteq> {}"
 | |
| 1209 | "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m" | |
| 1210 | "\<And>e. e>0 \<Longrightarrow> \<exists>n. \<forall>x\<in>S n. \<forall>y\<in>S n. dist x y < e" | |
| 1211 | obtains a where "\<And>n. a \<in> S n" | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1212 | proof - | 
| 68302 | 1213 | have "\<forall>n. \<exists>x. x \<in> S n" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 | using assms(2) by auto | 
| 68302 | 1215 | then have "\<exists>t. \<forall>n. t n \<in> S n" | 
| 1216 | using choice[of "\<lambda>n x. x \<in> S n"] by auto | |
| 1217 | then obtain t where t: "\<forall>n. t n \<in> S n" by auto | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | fix e :: real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1220 | assume "e > 0" | 
| 68302 | 1221 | then obtain N where N: "\<forall>x\<in>S N. \<forall>y\<in>S N. dist x y < e" | 
| 1222 | using assms(4) by blast | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1223 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 | fix m n :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1225 | assume "N \<le> m \<and> N \<le> n" | 
| 68302 | 1226 | then have "t m \<in> S N" "t n \<in> S N" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | using assms(3) t unfolding subset_eq t by blast+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1228 | then have "dist (t m) (t n) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | using N by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1230 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1231 | then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1232 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1233 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1234 | then have "Cauchy t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1235 | unfolding cauchy_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | then obtain l where l:"(t \<longlongrightarrow> l) sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | using complete_UNIV unfolding complete_def by auto | 
| 68302 | 1238 |   { fix n :: nat
 | 
| 1239 |     { fix e :: real
 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1240 | assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | then obtain N :: nat where N: "\<forall>n\<ge>N. dist (t n) l < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | using l[unfolded lim_sequentially] by auto | 
| 68302 | 1243 | have "t (max n N) \<in> S n" | 
| 66835 | 1244 | by (meson assms(3) contra_subsetD max.cobounded1 t) | 
| 68302 | 1245 | then have "\<exists>y\<in>S n. dist y l < e" | 
| 66835 | 1246 | using N max.cobounded2 by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1247 | } | 
| 68302 | 1248 | then have "l \<in> S n" | 
| 1249 | using closed_approachable[of "S n" l] assms(1) by auto | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1250 | } | 
| 68302 | 1251 | then show ?thesis | 
| 1252 | using that by blast | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1253 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1254 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | text \<open>Strengthen it to the intersection actually being a singleton.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 | lemma decreasing_closed_nest_sing: | 
| 68302 | 1258 | fixes S :: "nat \<Rightarrow> 'a::complete_space set" | 
| 1259 | assumes "\<And>n. closed(S n)" | |
| 1260 |           "\<And>n. S n \<noteq> {}"
 | |
| 1261 | "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m" | |
| 1262 | "\<And>e. e>0 \<Longrightarrow> \<exists>n. \<forall>x \<in> (S n). \<forall> y\<in>(S n). dist x y < e" | |
| 1263 |   shows "\<exists>a. \<Inter>(range S) = {a}"
 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | proof - | 
| 68302 | 1265 | obtain a where a: "\<forall>n. a \<in> S n" | 
| 1266 | using decreasing_closed_nest[of S] using assms by auto | |
| 1267 |   { fix b
 | |
| 1268 | assume b: "b \<in> \<Inter>(range S)" | |
| 1269 |     { fix e :: real
 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | then have "dist a b < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 | using assms(4) and b and a by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1274 | then have "dist a b = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | by (metis dist_eq_0_iff dist_nz less_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | } | 
| 68302 | 1277 |   with a have "\<Inter>(range S) = {a}"
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1278 | unfolding image_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1279 | then show ?thesis .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1281 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1283 | subsection \<open>Infimum Distance\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 | |
| 67962 | 1285 | definition%important "infdist x A = (if A = {} then 0 else INF a:A. dist x a)"
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1286 | |
| 67459 | 1287 | lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1288 | by (auto intro!: zero_le_dist) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1289 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1290 | lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a:A. dist x a)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1291 | by (simp add: infdist_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1292 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1293 | lemma infdist_nonneg: "0 \<le> infdist x A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 | by (auto simp: infdist_def intro: cINF_greatest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1295 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1296 | lemma infdist_le: "a \<in> A \<Longrightarrow> infdist x A \<le> dist x a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1297 | by (auto intro: cINF_lower simp add: infdist_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1298 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1299 | lemma infdist_le2: "a \<in> A \<Longrightarrow> dist x a \<le> d \<Longrightarrow> infdist x A \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1300 | by (auto intro!: cINF_lower2 simp add: infdist_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1301 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1302 | lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1303 | by (auto intro!: antisym infdist_nonneg infdist_le2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1304 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1305 | lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | proof (cases "A = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1308 | then show ?thesis by (simp add: infdist_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1309 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1310 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1311 | then obtain a where "a \<in> A" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1312 |   have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1313 | proof (rule cInf_greatest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1314 |     from \<open>A \<noteq> {}\<close> show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1315 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1316 | fix d | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1317 |     assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1318 | then obtain a where d: "d = dist x y + dist y a" "a \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1319 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1320 | show "infdist x A \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1321 |       unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>]
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1322 | proof (rule cINF_lower2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 | show "a \<in> A" by fact | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1324 | show "dist x a \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1325 | unfolding d by (rule dist_triangle) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1326 | qed simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1327 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1328 | also have "\<dots> = dist x y + infdist y A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1329 | proof (rule cInf_eq, safe) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | fix a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1331 | assume "a \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1332 | then show "dist x y + infdist y A \<le> dist x y + dist y a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1333 | by (auto intro: infdist_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1335 | fix i | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 |     assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1337 | then have "i - dist x y \<le> infdist y A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1338 |       unfolding infdist_notempty[OF \<open>A \<noteq> {}\<close>] using \<open>a \<in> A\<close>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | by (intro cINF_greatest) (auto simp: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1340 | then show "i \<le> dist x y + infdist y A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1343 | finally show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | lemma in_closure_iff_infdist_zero: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1347 |   assumes "A \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1348 | shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1349 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1350 | assume "x \<in> closure A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1351 | show "infdist x A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1352 | proof (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | assume "infdist x A \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | with infdist_nonneg[of x A] have "infdist x A > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1355 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 |     then have "ball x (infdist x A) \<inter> closure A = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | apply (metis \<open>x \<in> closure A\<close> closure_approachable dist_commute infdist_le not_less) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | then have "x \<notin> closure A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 | by (metis \<open>0 < infdist x A\<close> centre_in_ball disjoint_iff_not_equal) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | then show False using \<open>x \<in> closure A\<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1363 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1364 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 | assume x: "infdist x A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | then obtain a where "a \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | by atomize_elim (metis all_not_in_conv assms) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1368 | show "x \<in> closure A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1369 | unfolding closure_approachable | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1370 | apply safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1371 | proof (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1372 | fix e :: real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1373 | assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1374 | assume "\<not> (\<exists>y\<in>A. dist y x < e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | then have "infdist x A \<ge> e" using \<open>a \<in> A\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | unfolding infdist_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | by (force simp: dist_commute intro: cINF_greatest) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1378 | with x \<open>e > 0\<close> show False by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1379 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | lemma in_closed_iff_infdist_zero: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 |   assumes "closed A" "A \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | shows "x \<in> A \<longleftrightarrow> infdist x A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1386 | have "x \<in> closure A \<longleftrightarrow> infdist x A = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | by (rule in_closure_iff_infdist_zero) fact | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1388 | with assms show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1389 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | |
| 67455 | 1391 | lemma infdist_pos_not_in_closed: | 
| 1392 |   assumes "closed S" "S \<noteq> {}" "x \<notin> S"
 | |
| 1393 | shows "infdist x S > 0" | |
| 1394 | using in_closed_iff_infdist_zero[OF assms(1) assms(2), of x] assms(3) infdist_nonneg le_less by fastforce | |
| 1395 | ||
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1396 | lemma | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1397 | infdist_attains_inf: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1398 | fixes X::"'a::heine_borel set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1399 | assumes "closed X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1400 |   assumes "X \<noteq> {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1401 | obtains x where "x \<in> X" "infdist y X = dist y x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1402 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1403 | have "bdd_below (dist y ` X)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1404 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1405 | from distance_attains_inf[OF assms, of y] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1406 | obtain x where INF: "x \<in> X" "\<And>z. z \<in> X \<Longrightarrow> dist y x \<le> dist y z" by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1407 | have "infdist y X = dist y x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1408 | by (auto simp: infdist_def assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1409 | intro!: antisym cINF_lower[OF _ \<open>x \<in> X\<close>] cINF_greatest[OF assms(2) INF(2)]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1410 | with \<open>x \<in> X\<close> show ?thesis .. | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1411 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1412 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1413 | |
| 67455 | 1414 | text \<open>Every metric space is a T4 space:\<close> | 
| 1415 | ||
| 1416 | instance metric_space \<subseteq> t4_space | |
| 1417 | proof | |
| 1418 |   fix S T::"'a set" assume H: "closed S" "closed T" "S \<inter> T = {}"
 | |
| 1419 |   consider "S = {}" | "T = {}" | "S \<noteq> {} \<and> T \<noteq> {}" by auto
 | |
| 1420 |   then show "\<exists>U V. open U \<and> open V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> U \<inter> V = {}"
 | |
| 1421 | proof (cases) | |
| 1422 | case 1 | |
| 1423 | show ?thesis | |
| 1424 |       apply (rule exI[of _ "{}"], rule exI[of _ UNIV]) using 1 by auto
 | |
| 1425 | next | |
| 1426 | case 2 | |
| 1427 | show ?thesis | |
| 1428 |       apply (rule exI[of _ UNIV], rule exI[of _ "{}"]) using 2 by auto
 | |
| 1429 | next | |
| 1430 | case 3 | |
| 1431 | define U where "U = (\<Union>x\<in>S. ball x ((infdist x T)/2))" | |
| 1432 | have A: "open U" unfolding U_def by auto | |
| 1433 | have "infdist x T > 0" if "x \<in> S" for x | |
| 1434 | using H that 3 by (auto intro!: infdist_pos_not_in_closed) | |
| 1435 | then have B: "S \<subseteq> U" unfolding U_def by auto | |
| 1436 | define V where "V = (\<Union>x\<in>T. ball x ((infdist x S)/2))" | |
| 1437 | have C: "open V" unfolding V_def by auto | |
| 1438 | have "infdist x S > 0" if "x \<in> T" for x | |
| 1439 | using H that 3 by (auto intro!: infdist_pos_not_in_closed) | |
| 1440 | then have D: "T \<subseteq> V" unfolding V_def by auto | |
| 1441 | ||
| 1442 |     have "(ball x ((infdist x T)/2)) \<inter> (ball y ((infdist y S)/2)) = {}" if "x \<in> S" "y \<in> T" for x y
 | |
| 1443 | proof (auto) | |
| 1444 | fix z assume H: "dist x z * 2 < infdist x T" "dist y z * 2 < infdist y S" | |
| 1445 | have "2 * dist x y \<le> 2 * dist x z + 2 * dist y z" | |
| 1446 | using dist_triangle[of x y z] by (auto simp add: dist_commute) | |
| 1447 | also have "... < infdist x T + infdist y S" | |
| 1448 | using H by auto | |
| 1449 | finally have "dist x y < infdist x T \<or> dist x y < infdist y S" | |
| 1450 | by auto | |
| 1451 | then show False | |
| 1452 | using infdist_le[OF \<open>x \<in> S\<close>, of y] infdist_le[OF \<open>y \<in> T\<close>, of x] by (auto simp add: dist_commute) | |
| 1453 | qed | |
| 1454 |     then have E: "U \<inter> V = {}"
 | |
| 1455 | unfolding U_def V_def by auto | |
| 1456 | show ?thesis | |
| 1457 | apply (rule exI[of _ U], rule exI[of _ V]) using A B C D E by auto | |
| 1458 | qed | |
| 1459 | qed | |
| 1460 | ||
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1461 | lemma tendsto_infdist [tendsto_intros]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1462 | assumes f: "(f \<longlongrightarrow> l) F" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1463 | shows "((\<lambda>x. infdist (f x) A) \<longlongrightarrow> infdist l A) F" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1464 | proof (rule tendstoI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1465 | fix e ::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1466 | assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1467 | from tendstoD[OF f this] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1468 | show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | proof (eventually_elim) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1470 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1471 | from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1472 | have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1473 | by (simp add: dist_commute dist_real_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1474 | also assume "dist (f x) l < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1475 | finally show "dist (infdist (f x) A) (infdist l A) < e" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1476 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1477 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1478 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | lemma continuous_infdist[continuous_intros]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1480 | assumes "continuous F f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 | shows "continuous F (\<lambda>x. infdist (f x) A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1482 | using assms unfolding continuous_def by (rule tendsto_infdist) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1484 | lemma compact_infdist_le: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1485 | fixes A::"'a::heine_borel set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1486 |   assumes "A \<noteq> {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1487 | assumes "compact A" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1488 | assumes "e > 0" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1489 |   shows "compact {x. infdist x A \<le> e}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1490 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1491 |   from continuous_closed_vimage[of "{0..e}" "\<lambda>x. infdist x A"]
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1492 | continuous_infdist[OF continuous_ident, of _ UNIV A] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1493 |   have "closed {x. infdist x A \<le> e}" by (auto simp: vimage_def infdist_nonneg)
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1494 | moreover | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1495 | from assms obtain x0 b where b: "\<And>x. x \<in> A \<Longrightarrow> dist x0 x \<le> b" "closed A" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1496 | by (auto simp: compact_eq_bounded_closed bounded_def) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1497 |   {
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1498 | fix y | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1499 | assume le: "infdist y A \<le> e" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1500 |     from infdist_attains_inf[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>, of y]
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1501 | obtain z where z: "z \<in> A" "infdist y A = dist y z" by blast | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1502 | have "dist x0 y \<le> dist y z + dist x0 z" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1503 | by (metis dist_commute dist_triangle) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1504 | also have "dist y z \<le> e" using le z by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1505 | also have "dist x0 z \<le> b" using b z by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1506 | finally have "dist x0 y \<le> b + e" by arith | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1507 | } then | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1508 |   have "bounded {x. infdist x A \<le> e}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1509 | by (auto simp: bounded_any_center[where a=x0] intro!: exI[where x="b + e"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1510 |   ultimately show "compact {x. infdist x A \<le> e}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1511 | by (simp add: compact_eq_bounded_closed) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1512 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1513 | |
| 67968 | 1514 | subsection%unimportant \<open>Equality of continuous functions on closure and related results\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1515 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1516 | lemma continuous_closedin_preimage_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1517 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1518 |   shows "continuous_on S f \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x = a}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1519 |   using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq)
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1520 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1521 | lemma continuous_closed_preimage_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1522 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1523 |   shows "continuous_on S f \<Longrightarrow> closed S \<Longrightarrow> closed {x \<in> S. f x = a}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1524 |   using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq)
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1525 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1526 | lemma continuous_constant_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1527 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1528 | assumes "continuous_on (closure S) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1529 | and "\<And>x. x \<in> S \<Longrightarrow> f x = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1530 | and "x \<in> closure S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1531 | shows "f x = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1532 | using continuous_closed_preimage_constant[of "closure S" f a] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1533 |       assms closure_minimal[of S "{x \<in> closure S. f x = a}"] closure_subset
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1534 | unfolding subset_eq | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1535 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1536 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1537 | lemma image_closure_subset: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1538 | assumes contf: "continuous_on (closure S) f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1539 | and "closed T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1540 | and "(f ` S) \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1541 | shows "f ` (closure S) \<subseteq> T" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1542 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1543 |   have "S \<subseteq> {x \<in> closure S. f x \<in> T}"
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1544 | using assms(3) closure_subset by auto | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1545 | moreover have "closed (closure S \<inter> f -` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1546 | using continuous_closed_preimage[OF contf] \<open>closed T\<close> by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1547 | ultimately have "closure S = (closure S \<inter> f -` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 1548 | using closure_minimal[of S "(closure S \<inter> f -` T)"] by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1549 | then show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1550 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1551 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1552 | lemma continuous_on_closure_norm_le: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1554 | assumes "continuous_on (closure s) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1555 | and "\<forall>y \<in> s. norm(f y) \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1556 | and "x \<in> (closure s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | shows "norm (f x) \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1558 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1559 | have *: "f ` s \<subseteq> cball 0 b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | using assms(2)[unfolded mem_cball_0[symmetric]] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1561 | show ?thesis | 
| 66835 | 1562 | by (meson "*" assms(1) assms(3) closed_cball image_closure_subset image_subset_iff mem_cball_0) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1563 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1564 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1565 | lemma isCont_indicator: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1566 | fixes x :: "'a::t2_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1567 | shows "isCont (indicator A :: 'a \<Rightarrow> real) x = (x \<notin> frontier A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1568 | proof auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1569 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1570 | assume cts_at: "isCont (indicator A :: 'a \<Rightarrow> real) x" and fr: "x \<in> frontier A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1571 | with continuous_at_open have 1: "\<forall>V::real set. open V \<and> indicator A x \<in> V \<longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1572 | (\<exists>U::'a set. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> V))" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1573 | show False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | proof (cases "x \<in> A") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1575 | assume x: "x \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1576 |     hence "indicator A x \<in> ({0<..<2} :: real set)" by simp
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 |     hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({0<..<2} :: real set))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1578 | using 1 open_greaterThanLessThan by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | then guess U .. note U = this | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1580 | hence "\<forall>y\<in>U. indicator A y > (0::real)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1581 | unfolding greaterThanLessThan_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1582 | hence "U \<subseteq> A" using indicator_eq_0_iff by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1583 | hence "x \<in> interior A" using U interiorI by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1584 | thus ?thesis using fr unfolding frontier_def by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1585 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1586 | assume x: "x \<notin> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1587 |     hence "indicator A x \<in> ({-1<..<1} :: real set)" by simp
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1588 |     hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({-1<..<1} :: real set))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1589 | using 1 open_greaterThanLessThan by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1590 | then guess U .. note U = this | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1591 | hence "\<forall>y\<in>U. indicator A y < (1::real)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1592 | unfolding greaterThanLessThan_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1593 | hence "U \<subseteq> -A" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1594 | hence "x \<in> interior (-A)" using U interiorI by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1595 | thus ?thesis using fr interior_complement unfolding frontier_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1596 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1597 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | assume nfr: "x \<notin> frontier A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1599 | hence "x \<in> interior A \<or> x \<in> interior (-A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1600 | by (auto simp: frontier_def closure_interior) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1601 | thus "isCont ((indicator A)::'a \<Rightarrow> real) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1602 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1603 | assume int: "x \<in> interior A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1604 | then obtain U where U: "open U" "x \<in> U" "U \<subseteq> A" unfolding interior_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1605 | hence "\<forall>y\<in>U. indicator A y = (1::real)" unfolding indicator_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1606 | hence "continuous_on U (indicator A)" by (simp add: continuous_on_const indicator_eq_1_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1607 | thus ?thesis using U continuous_on_eq_continuous_at by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1608 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1609 | assume ext: "x \<in> interior (-A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1610 | then obtain U where U: "open U" "x \<in> U" "U \<subseteq> -A" unfolding interior_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1611 | then have "continuous_on U (indicator A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1612 | using continuous_on_topological by (auto simp: subset_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1613 | thus ?thesis using U continuous_on_eq_continuous_at by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1614 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1615 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1616 | |
| 67962 | 1617 | subsection%unimportant \<open>A function constant on a set\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1618 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1619 | definition constant_on (infixl "(constant'_on)" 50) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1620 | where "f constant_on A \<equiv> \<exists>y. \<forall>x\<in>A. f x = y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1621 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1622 | lemma constant_on_subset: "\<lbrakk>f constant_on A; B \<subseteq> A\<rbrakk> \<Longrightarrow> f constant_on B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1623 | unfolding constant_on_def by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1624 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1625 | lemma injective_not_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1626 |   fixes S :: "'a::{perfect_space} set"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1627 |   shows "\<lbrakk>open S; inj_on f S; f constant_on S\<rbrakk> \<Longrightarrow> S = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1628 | unfolding constant_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1629 | by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1630 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1631 | lemma constant_on_closureI: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1632 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1633 | assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1634 | shows "f constant_on (closure S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1635 | using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1636 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1637 | |
| 67962 | 1638 | subsection%unimportant\<open>Relating linear images to open/closed/interior/closure\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1639 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1640 | proposition open_surjective_linear_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1641 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1642 | assumes "open A" "linear f" "surj f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1643 | shows "open(f ` A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1644 | unfolding open_dist | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1645 | proof clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1646 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1647 | assume "x \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1648 | have "bounded (inv f ` Basis)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1649 | by (simp add: finite_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1650 | with bounded_pos obtain B where "B > 0" and B: "\<And>x. x \<in> inv f ` Basis \<Longrightarrow> norm x \<le> B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1651 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1652 | obtain e where "e > 0" and e: "\<And>z. dist z x < e \<Longrightarrow> z \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1653 | by (metis open_dist \<open>x \<in> A\<close> \<open>open A\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1654 |   define \<delta> where "\<delta> \<equiv> e / B / DIM('b)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1655 | show "\<exists>e>0. \<forall>y. dist y (f x) < e \<longrightarrow> y \<in> f ` A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1656 | proof (intro exI conjI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1657 | show "\<delta> > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1658 | using \<open>e > 0\<close> \<open>B > 0\<close> by (simp add: \<delta>_def divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1659 |     have "y \<in> f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1660 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1661 | define u where "u \<equiv> y - f x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1662 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1663 | proof (rule image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1664 | show "y = f (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1665 | apply (simp add: linear_add linear_sum linear.scaleR \<open>linear f\<close> surj_f_inv_f \<open>surj f\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1666 | apply (simp add: euclidean_representation u_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1667 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1668 | have "dist (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i)) x \<le> (\<Sum>i\<in>Basis. norm ((u \<bullet> i) *\<^sub>R inv f i))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1669 | by (simp add: dist_norm sum_norm_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1670 | also have "... = (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar> * norm (inv f i))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1671 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1672 | also have "... \<le> (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar>) * B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1673 | by (simp add: B sum_distrib_right sum_mono mult_left_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1674 |         also have "... \<le> DIM('b) * dist y (f x) * B"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1675 | apply (rule mult_right_mono [OF sum_bounded_above]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1676 | using \<open>0 < B\<close> by (auto simp: Basis_le_norm dist_norm u_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1677 | also have "... < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1678 | by (metis mult.commute mult.left_commute that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1679 | finally show "x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i) \<in> A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1680 | by (rule e) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1681 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1682 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1683 | then show "\<forall>y. dist y (f x) < \<delta> \<longrightarrow> y \<in> f ` A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1684 | using \<open>e > 0\<close> \<open>B > 0\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1685 | by (auto simp: \<delta>_def divide_simps mult_less_0_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1686 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1687 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1688 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1689 | corollary open_bijective_linear_image_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1690 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1691 | assumes "linear f" "bij f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1692 | shows "open(f ` A) \<longleftrightarrow> open A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1693 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1694 | assume "open(f ` A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1695 | then have "open(f -` (f ` A))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1696 | using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1697 | then show "open A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1698 | by (simp add: assms bij_is_inj inj_vimage_image_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1699 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1700 | assume "open A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1701 | then show "open(f ` A)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1702 | by (simp add: assms bij_is_surj open_surjective_linear_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1703 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1704 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1705 | corollary interior_bijective_linear_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1706 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1707 | assumes "linear f" "bij f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1708 | shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1709 | proof safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1710 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1711 | assume x: "x \<in> ?lhs" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1712 | then obtain T where "open T" and "x \<in> T" and "T \<subseteq> f ` S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1713 | by (metis interiorE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1714 | then show "x \<in> ?rhs" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1715 | by (metis (no_types, hide_lams) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1716 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1717 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1718 | assume x: "x \<in> interior S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1719 | then show "f x \<in> interior (f ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1720 | by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1721 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1722 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1723 | lemma interior_injective_linear_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1724 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1725 | assumes "linear f" "inj f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1726 | shows "interior(f ` S) = f ` (interior S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1727 | by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1728 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1729 | lemma interior_surjective_linear_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1730 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1731 | assumes "linear f" "surj f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1732 | shows "interior(f ` S) = f ` (interior S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1733 | by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1734 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1735 | lemma interior_negations: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1736 | fixes S :: "'a::euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1737 | shows "interior(uminus ` S) = image uminus (interior S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1738 | by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1739 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1740 | text \<open>Preservation of compactness and connectedness under continuous function.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1741 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1742 | lemma compact_eq_openin_cover: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1743 | "compact S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1744 | (\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1745 | (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1746 | proof safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1747 | fix C | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1748 | assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1749 |   then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1750 | unfolding openin_open by force+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1751 |   with \<open>compact S\<close> obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1752 | by (meson compactE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1753 | then have "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1754 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1755 | then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1756 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1757 | assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1758 | (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1759 | show "compact S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1760 | proof (rule compactI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1761 | fix C | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1762 | let ?C = "image (\<lambda>T. S \<inter> T) C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1763 | assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1764 | then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1765 | unfolding openin_open by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1766 | with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1767 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1768 | let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1769 | have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1770 | proof (intro conjI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1771 | from \<open>D \<subseteq> ?C\<close> show "?D \<subseteq> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1772 | by (fast intro: inv_into_into) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1773 | from \<open>finite D\<close> show "finite ?D" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1774 | by (rule finite_imageI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1775 | from \<open>S \<subseteq> \<Union>D\<close> show "S \<subseteq> \<Union>?D" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1776 | apply (rule subset_trans, clarsimp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1777 | apply (frule subsetD [OF \<open>D \<subseteq> ?C\<close>, THEN f_inv_into_f]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1778 | apply (erule rev_bexI, fast) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1779 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1780 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1781 | then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1782 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1783 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1784 | |
| 67962 | 1785 | subsection%unimportant\<open> Theorems relating continuity and uniform continuity to closures\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1786 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1787 | lemma continuous_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1788 | "continuous_on (closure S) f \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1789 | (\<forall>x e. x \<in> closure S \<and> 0 < e | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1790 | \<longrightarrow> (\<exists>d. 0 < d \<and> (\<forall>y. y \<in> S \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1791 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1792 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1793 | assume ?lhs then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1794 | unfolding continuous_on_iff by (metis Un_iff closure_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1795 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1796 | assume R [rule_format]: ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1797 | show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1798 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1799 | fix x and e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1800 | assume "0 < e" and x: "x \<in> closure S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1801 | obtain \<delta>::real where "\<delta> > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1802 | and \<delta>: "\<And>y. \<lbrakk>y \<in> S; dist y x < \<delta>\<rbrakk> \<Longrightarrow> dist (f y) (f x) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1803 | using R [of x "e/2"] \<open>0 < e\<close> x by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1804 | have "dist (f y) (f x) \<le> e" if y: "y \<in> closure S" and dyx: "dist y x < \<delta>/2" for y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1805 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1806 | obtain \<delta>'::real where "\<delta>' > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1807 | and \<delta>': "\<And>z. \<lbrakk>z \<in> S; dist z y < \<delta>'\<rbrakk> \<Longrightarrow> dist (f z) (f y) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1808 | using R [of y "e/2"] \<open>0 < e\<close> y by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1809 | obtain z where "z \<in> S" and z: "dist z y < min \<delta>' \<delta> / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1810 | using closure_approachable y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1811 | by (metis \<open>0 < \<delta>'\<close> \<open>0 < \<delta>\<close> divide_pos_pos min_less_iff_conj zero_less_numeral) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1812 | have "dist (f z) (f y) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1813 | apply (rule \<delta>' [OF \<open>z \<in> S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1814 | using z \<open>0 < \<delta>'\<close> by linarith | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1815 | moreover have "dist (f z) (f x) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1816 | apply (rule \<delta> [OF \<open>z \<in> S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1817 | using z \<open>0 < \<delta>\<close> dist_commute[of y z] dist_triangle_half_r [of y] dyx by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1818 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1819 | by (metis dist_commute dist_triangle_half_l less_imp_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1820 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1821 | then show "\<exists>d>0. \<forall>x'\<in>closure S. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1822 | by (rule_tac x="\<delta>/2" in exI) (simp add: \<open>\<delta> > 0\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1823 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1824 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1825 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1826 | lemma continuous_on_closure_sequentially: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1827 | fixes f :: "'a::metric_space \<Rightarrow> 'b :: metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1828 | shows | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1829 | "continuous_on (closure S) f \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1830 | (\<forall>x a. a \<in> closure S \<and> (\<forall>n. x n \<in> S) \<and> x \<longlonglongrightarrow> a \<longrightarrow> (f \<circ> x) \<longlonglongrightarrow> f a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1831 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1832 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1833 | have "continuous_on (closure S) f \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1834 | (\<forall>x \<in> closure S. continuous (at x within S) f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1835 | by (force simp: continuous_on_closure continuous_within_eps_delta) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1836 | also have "... = ?rhs" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1837 | by (force simp: continuous_within_sequentially) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1838 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1839 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1840 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1841 | lemma uniformly_continuous_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1842 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1843 | assumes ucont: "uniformly_continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1844 | and cont: "continuous_on (closure S) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1845 | shows "uniformly_continuous_on (closure S) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1846 | unfolding uniformly_continuous_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1847 | proof (intro allI impI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1848 | fix e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1849 | assume "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1850 | then obtain d::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1851 | where "d>0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1852 | and d: "\<And>x x'. \<lbrakk>x\<in>S; x'\<in>S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1853 | using ucont [unfolded uniformly_continuous_on_def, rule_format, of "e/3"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1854 | show "\<exists>d>0. \<forall>x\<in>closure S. \<forall>x'\<in>closure S. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1855 | proof (rule exI [where x="d/3"], clarsimp simp: \<open>d > 0\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1856 | fix x y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1857 | assume x: "x \<in> closure S" and y: "y \<in> closure S" and dyx: "dist y x * 3 < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1858 | obtain d1::real where "d1 > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1859 | and d1: "\<And>w. \<lbrakk>w \<in> closure S; dist w x < d1\<rbrakk> \<Longrightarrow> dist (f w) (f x) < e/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1860 | using cont [unfolded continuous_on_iff, rule_format, of "x" "e/3"] \<open>0 < e\<close> x by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1861 | obtain x' where "x' \<in> S" and x': "dist x' x < min d1 (d / 3)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1862 | using closure_approachable [of x S] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1863 | by (metis \<open>0 < d1\<close> \<open>0 < d\<close> divide_pos_pos min_less_iff_conj x zero_less_numeral) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1864 | obtain d2::real where "d2 > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1865 | and d2: "\<forall>w \<in> closure S. dist w y < d2 \<longrightarrow> dist (f w) (f y) < e/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1866 | using cont [unfolded continuous_on_iff, rule_format, of "y" "e/3"] \<open>0 < e\<close> y by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1867 | obtain y' where "y' \<in> S" and y': "dist y' y < min d2 (d / 3)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1868 | using closure_approachable [of y S] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1869 | by (metis \<open>0 < d2\<close> \<open>0 < d\<close> divide_pos_pos min_less_iff_conj y zero_less_numeral) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1870 | have "dist x' x < d/3" using x' by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1871 | moreover have "dist x y < d/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1872 | by (metis dist_commute dyx less_divide_eq_numeral1(1)) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1873 | moreover have "dist y y' < d/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1874 | by (metis (no_types) dist_commute min_less_iff_conj y') | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1875 | ultimately have "dist x' y' < d/3 + d/3 + d/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1876 | by (meson dist_commute_lessI dist_triangle_lt add_strict_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1877 | then have "dist x' y' < d" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1878 | then have "dist (f x') (f y') < e/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1879 | by (rule d [OF \<open>y' \<in> S\<close> \<open>x' \<in> S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1880 | moreover have "dist (f x') (f x) < e/3" using \<open>x' \<in> S\<close> closure_subset x' d1 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1881 | by (simp add: closure_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1882 | moreover have "dist (f y') (f y) < e/3" using \<open>y' \<in> S\<close> closure_subset y' d2 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1883 | by (simp add: closure_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1884 | ultimately have "dist (f y) (f x) < e/3 + e/3 + e/3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1885 | by (meson dist_commute_lessI dist_triangle_lt add_strict_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1886 | then show "dist (f y) (f x) < e" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1887 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1888 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1889 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1890 | lemma uniformly_continuous_on_extension_at_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1891 | fixes f::"'a::metric_space \<Rightarrow> 'b::complete_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1892 | assumes uc: "uniformly_continuous_on X f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1893 | assumes "x \<in> closure X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1894 | obtains l where "(f \<longlongrightarrow> l) (at x within X)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1895 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1896 | from assms obtain xs where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1897 | by (auto simp: closure_sequential) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1898 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1899 | from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF xs] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1900 | obtain l where l: "(\<lambda>n. f (xs n)) \<longlonglongrightarrow> l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1901 | by atomize_elim (simp only: convergent_eq_Cauchy) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1902 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1903 | have "(f \<longlongrightarrow> l) (at x within X)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1904 | proof (safe intro!: Lim_within_LIMSEQ) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1905 | fix xs' | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1906 | assume "\<forall>n. xs' n \<noteq> x \<and> xs' n \<in> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1907 | and xs': "xs' \<longlonglongrightarrow> x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1908 | then have "xs' n \<noteq> x" "xs' n \<in> X" for n by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1909 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1910 | from uniformly_continuous_on_Cauchy[OF uc LIMSEQ_imp_Cauchy, OF \<open>xs' \<longlonglongrightarrow> x\<close> \<open>xs' _ \<in> X\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1911 | obtain l' where l': "(\<lambda>n. f (xs' n)) \<longlonglongrightarrow> l'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1912 | by atomize_elim (simp only: convergent_eq_Cauchy) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1913 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1914 | show "(\<lambda>n. f (xs' n)) \<longlonglongrightarrow> l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1915 | proof (rule tendstoI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1916 | fix e::real assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1917 | define e' where "e' \<equiv> e / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1918 | have "e' > 0" using \<open>e > 0\<close> by (simp add: e'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1919 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1920 | have "\<forall>\<^sub>F n in sequentially. dist (f (xs n)) l < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1921 | by (simp add: \<open>0 < e'\<close> l tendstoD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1922 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1923 | from uc[unfolded uniformly_continuous_on_def, rule_format, OF \<open>e' > 0\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1924 | obtain d where d: "d > 0" "\<And>x x'. x \<in> X \<Longrightarrow> x' \<in> X \<Longrightarrow> dist x x' < d \<Longrightarrow> dist (f x) (f x') < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1925 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1926 | have "\<forall>\<^sub>F n in sequentially. dist (xs n) (xs' n) < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1927 | by (auto intro!: \<open>0 < d\<close> order_tendstoD tendsto_eq_intros xs xs') | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1928 | ultimately | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1929 | show "\<forall>\<^sub>F n in sequentially. dist (f (xs' n)) l < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1930 | proof eventually_elim | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1931 | case (elim n) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1932 | have "dist (f (xs' n)) l \<le> dist (f (xs n)) (f (xs' n)) + dist (f (xs n)) l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1933 | by (metis dist_triangle dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1934 | also have "dist (f (xs n)) (f (xs' n)) < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1935 | by (auto intro!: d xs \<open>xs' _ \<in> _\<close> elim) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1936 | also note \<open>dist (f (xs n)) l < e'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1937 | also have "e' + e' = e" by (simp add: e'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1938 | finally show ?case by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1939 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1940 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1941 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1942 | thus ?thesis .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1943 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1944 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1945 | lemma uniformly_continuous_on_extension_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1946 | fixes f::"'a::metric_space \<Rightarrow> 'b::complete_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1947 | assumes uc: "uniformly_continuous_on X f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1948 | obtains g where "uniformly_continuous_on (closure X) g" "\<And>x. x \<in> X \<Longrightarrow> f x = g x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1949 | "\<And>Y h x. X \<subseteq> Y \<Longrightarrow> Y \<subseteq> closure X \<Longrightarrow> continuous_on Y h \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> f x = h x) \<Longrightarrow> x \<in> Y \<Longrightarrow> h x = g x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1950 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1951 | from uc have cont_f: "continuous_on X f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1952 | by (simp add: uniformly_continuous_imp_continuous) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1953 | obtain y where y: "(f \<longlongrightarrow> y x) (at x within X)" if "x \<in> closure X" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1954 | apply atomize_elim | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1955 | apply (rule choice) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1956 | using uniformly_continuous_on_extension_at_closure[OF assms] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1957 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1958 | let ?g = "\<lambda>x. if x \<in> X then f x else y x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1959 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1960 | have "uniformly_continuous_on (closure X) ?g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1961 | unfolding uniformly_continuous_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1962 | proof safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1963 | fix e::real assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1964 | define e' where "e' \<equiv> e / 3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1965 | have "e' > 0" using \<open>e > 0\<close> by (simp add: e'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1966 | from uc[unfolded uniformly_continuous_on_def, rule_format, OF \<open>0 < e'\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1967 | obtain d where "d > 0" and d: "\<And>x x'. x \<in> X \<Longrightarrow> x' \<in> X \<Longrightarrow> dist x' x < d \<Longrightarrow> dist (f x') (f x) < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1968 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1969 | define d' where "d' = d / 3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1970 | have "d' > 0" using \<open>d > 0\<close> by (simp add: d'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1971 | show "\<exists>d>0. \<forall>x\<in>closure X. \<forall>x'\<in>closure X. dist x' x < d \<longrightarrow> dist (?g x') (?g x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1972 | proof (safe intro!: exI[where x=d'] \<open>d' > 0\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1973 | fix x x' assume x: "x \<in> closure X" and x': "x' \<in> closure X" and dist: "dist x' x < d'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1974 | then obtain xs xs' where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1975 | and xs': "xs' \<longlonglongrightarrow> x'" "\<And>n. xs' n \<in> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1976 | by (auto simp: closure_sequential) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1977 | have "\<forall>\<^sub>F n in sequentially. dist (xs' n) x' < d'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1978 | and "\<forall>\<^sub>F n in sequentially. dist (xs n) x < d'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1979 | by (auto intro!: \<open>0 < d'\<close> order_tendstoD tendsto_eq_intros xs xs') | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1980 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1981 | have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> y x" if "x \<in> closure X" "x \<notin> X" "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X" for xs x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1982 | using that not_eventuallyD | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1983 | by (force intro!: filterlim_compose[OF y[OF \<open>x \<in> closure X\<close>]] simp: filterlim_at) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1984 | then have "(\<lambda>x. f (xs' x)) \<longlonglongrightarrow> ?g x'" "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> ?g x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1985 | using x x' | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1986 | by (auto intro!: continuous_on_tendsto_compose[OF cont_f] simp: xs' xs) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1987 | then have "\<forall>\<^sub>F n in sequentially. dist (f (xs' n)) (?g x') < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1988 | "\<forall>\<^sub>F n in sequentially. dist (f (xs n)) (?g x) < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1989 | by (auto intro!: \<open>0 < e'\<close> order_tendstoD tendsto_eq_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1990 | ultimately | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1991 | have "\<forall>\<^sub>F n in sequentially. dist (?g x') (?g x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1992 | proof eventually_elim | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1993 | case (elim n) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1994 | have "dist (?g x') (?g x) \<le> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1995 | dist (f (xs' n)) (?g x') + dist (f (xs' n)) (f (xs n)) + dist (f (xs n)) (?g x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1996 | by (metis add.commute add_le_cancel_left dist_commute dist_triangle dist_triangle_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1997 | also | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1998 |         {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1999 | have "dist (xs' n) (xs n) \<le> dist (xs' n) x' + dist x' x + dist (xs n) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2000 | by (metis add.commute add_le_cancel_left dist_triangle dist_triangle_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2001 | also note \<open>dist (xs' n) x' < d'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2002 | also note \<open>dist x' x < d'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2003 | also note \<open>dist (xs n) x < d'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2004 | finally have "dist (xs' n) (xs n) < d" by (simp add: d'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2005 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2006 | with \<open>xs _ \<in> X\<close> \<open>xs' _ \<in> X\<close> have "dist (f (xs' n)) (f (xs n)) < e'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2007 | by (rule d) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2008 | also note \<open>dist (f (xs' n)) (?g x') < e'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2009 | also note \<open>dist (f (xs n)) (?g x) < e'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2010 | finally show ?case by (simp add: e'_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2011 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2012 | then show "dist (?g x') (?g x) < e" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2013 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2014 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2015 | moreover have "f x = ?g x" if "x \<in> X" for x using that by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2016 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2017 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2018 | fix Y h x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2019 | assume Y: "x \<in> Y" "X \<subseteq> Y" "Y \<subseteq> closure X" and cont_h: "continuous_on Y h" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2020 | and extension: "(\<And>x. x \<in> X \<Longrightarrow> f x = h x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2021 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2022 | assume "x \<notin> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2023 | have "x \<in> closure X" using Y by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2024 | then obtain xs where xs: "xs \<longlonglongrightarrow> x" "\<And>n. xs n \<in> X" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2025 | by (auto simp: closure_sequential) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2026 | from continuous_on_tendsto_compose[OF cont_h xs(1)] xs(2) Y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2027 | have hx: "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> h x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2028 | by (auto simp: set_mp extension) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2029 | then have "(\<lambda>x. f (xs x)) \<longlonglongrightarrow> y x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2030 | using \<open>x \<notin> X\<close> not_eventuallyD xs(2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2031 | by (force intro!: filterlim_compose[OF y[OF \<open>x \<in> closure X\<close>]] simp: filterlim_at xs) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2032 | with hx have "h x = y x" by (rule LIMSEQ_unique) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2033 | } then | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2034 | have "h x = ?g x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2035 | using extension by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2036 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2037 | ultimately show ?thesis .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2038 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2039 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2040 | lemma bounded_uniformly_continuous_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2041 | fixes f :: "'a :: heine_borel \<Rightarrow> 'b :: heine_borel" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2042 | assumes "uniformly_continuous_on S f" "bounded S" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2043 | shows "bounded(f ` S)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2044 | by (metis (no_types, lifting) assms bounded_closure_image compact_closure compact_continuous_image compact_eq_bounded_closed image_cong uniformly_continuous_imp_continuous uniformly_continuous_on_extension_on_closure) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2045 | |
| 67968 | 2046 | subsection%unimportant \<open>Making a continuous function avoid some value in a neighbourhood\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2047 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2048 | lemma continuous_within_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2049 | fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2050 | assumes "continuous (at x within s) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2051 | and "f x \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2052 | shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2053 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2054 | obtain U where "open U" and "f x \<in> U" and "a \<notin> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2055 | using t1_space [OF \<open>f x \<noteq> a\<close>] by fast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2056 | have "(f \<longlongrightarrow> f x) (at x within s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2057 | using assms(1) by (simp add: continuous_within) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2058 | then have "eventually (\<lambda>y. f y \<in> U) (at x within s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2059 | using \<open>open U\<close> and \<open>f x \<in> U\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2060 | unfolding tendsto_def by fast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2061 | then have "eventually (\<lambda>y. f y \<noteq> a) (at x within s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2062 | using \<open>a \<notin> U\<close> by (fast elim: eventually_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2063 | then show ?thesis | 
| 66953 | 2064 | using \<open>f x \<noteq> a\<close> by (auto simp: dist_commute eventually_at) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2065 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2066 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2067 | lemma continuous_at_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2068 | fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2069 | assumes "continuous (at x) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2070 | and "f x \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2071 | shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2072 | using assms continuous_within_avoid[of x UNIV f a] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2073 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2074 | lemma continuous_on_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2075 | fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2076 | assumes "continuous_on s f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2077 | and "x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2078 | and "f x \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2079 | shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2080 | using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2081 | OF assms(2)] continuous_within_avoid[of x s f a] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2082 | using assms(3) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2083 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2084 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2085 | lemma continuous_on_open_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2086 | fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2087 | assumes "continuous_on s f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2088 | and "open s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2089 | and "x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2090 | and "f x \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2091 | shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2092 | using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2093 | using continuous_at_avoid[of x f a] assms(4) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2094 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2095 | |
| 67962 | 2096 | subsection%unimportant\<open>Quotient maps\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2097 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2098 | lemma quotient_map_imp_continuous_open: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2099 | assumes T: "f ` S \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2100 | and ope: "\<And>U. U \<subseteq> T | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2101 | \<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2102 | openin (subtopology euclidean T) U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2103 | shows "continuous_on S f" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2104 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2105 | have [simp]: "S \<inter> f -` f ` S = S" by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2106 | show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2107 | using ope [OF T] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2108 | apply (simp add: continuous_on_open) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2109 | by (meson ope openin_imp_subset openin_trans) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2110 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2111 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2112 | lemma quotient_map_imp_continuous_closed: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2113 | assumes T: "f ` S \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2114 | and ope: "\<And>U. U \<subseteq> T | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2115 | \<Longrightarrow> (closedin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2116 | closedin (subtopology euclidean T) U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2117 | shows "continuous_on S f" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2118 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2119 | have [simp]: "S \<inter> f -` f ` S = S" by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2120 | show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2121 | using ope [OF T] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2122 | apply (simp add: continuous_on_closed) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2123 | by (metis (no_types, lifting) ope closedin_imp_subset closedin_trans) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2124 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2125 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2126 | lemma open_map_imp_quotient_map: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2127 | assumes contf: "continuous_on S f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2128 | and T: "T \<subseteq> f ` S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2129 | and ope: "\<And>T. openin (subtopology euclidean S) T | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2130 | \<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2131 | shows "openin (subtopology euclidean S) (S \<inter> f -` T) = | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2132 | openin (subtopology euclidean (f ` S)) T" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2133 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2134 | have "T = f ` (S \<inter> f -` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2135 | using T by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2136 | then show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2137 | using "ope" contf continuous_on_open by metis | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2138 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2139 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2140 | lemma closed_map_imp_quotient_map: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2141 | assumes contf: "continuous_on S f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2142 | and T: "T \<subseteq> f ` S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2143 | and ope: "\<And>T. closedin (subtopology euclidean S) T | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2144 | \<Longrightarrow> closedin (subtopology euclidean (f ` S)) (f ` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2145 | shows "openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2146 | openin (subtopology euclidean (f ` S)) T" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2147 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2148 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2149 | assume ?lhs | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2150 | then have *: "closedin (subtopology euclidean S) (S - (S \<inter> f -` T))" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2151 | using closedin_diff by fastforce | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2152 | have [simp]: "(f ` S - f ` (S - (S \<inter> f -` T))) = T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2153 | using T by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2154 | show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2155 | using ope [OF *, unfolded closedin_def] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2156 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2157 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2158 | with contf show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2159 | by (auto simp: continuous_on_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2160 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2161 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2162 | lemma continuous_right_inverse_imp_quotient_map: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2163 | assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2164 | and contg: "continuous_on T g" and img: "g ` T \<subseteq> S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2165 | and fg [simp]: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2166 | and U: "U \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2167 | shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2168 | openin (subtopology euclidean T) U" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2169 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2170 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2171 | have f: "\<And>Z. openin (subtopology euclidean (f ` S)) Z \<Longrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2172 | openin (subtopology euclidean S) (S \<inter> f -` Z)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2173 | and g: "\<And>Z. openin (subtopology euclidean (g ` T)) Z \<Longrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2174 | openin (subtopology euclidean T) (T \<inter> g -` Z)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2175 | using contf contg by (auto simp: continuous_on_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2176 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2177 | proof | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2178 |     have "T \<inter> g -` (g ` T \<inter> (S \<inter> f -` U)) = {x \<in> T. f (g x) \<in> U}"
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2179 | using imf img by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2180 | also have "... = U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2181 | using U by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2182 | finally have eq: "T \<inter> g -` (g ` T \<inter> (S \<inter> f -` U)) = U" . | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2183 | assume ?lhs | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2184 | then have *: "openin (subtopology euclidean (g ` T)) (g ` T \<inter> (S \<inter> f -` U))" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2185 | by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2186 | show ?rhs | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2187 | using g [OF *] eq by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2188 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2189 | assume rhs: ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2190 | show ?lhs | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2191 | by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2192 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2193 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2194 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2195 | lemma continuous_left_inverse_imp_quotient_map: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2196 | assumes "continuous_on S f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2197 | and "continuous_on (f ` S) g" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2198 | and "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2199 | and "U \<subseteq> f ` S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2200 | shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2201 | openin (subtopology euclidean (f ` S)) U" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2202 | apply (rule continuous_right_inverse_imp_quotient_map) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2203 | using assms apply force+ | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2204 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2205 | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 2206 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2207 | text \<open>Proving a function is constant by proving that a level set is open\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2208 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2209 | lemma continuous_levelset_openin_cases: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2210 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2211 | shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2212 |         openin (subtopology euclidean s) {x \<in> s. f x = a}
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2213 | \<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2214 | unfolding connected_clopen | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2215 | using continuous_closedin_preimage_constant by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2216 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2217 | lemma continuous_levelset_openin: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2218 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2219 | shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2220 |         openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2221 | (\<exists>x \<in> s. f x = a) \<Longrightarrow> (\<forall>x \<in> s. f x = a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2222 | using continuous_levelset_openin_cases[of s f ] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2223 | by meson | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2224 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2225 | lemma continuous_levelset_open: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2226 | fixes f :: "_ \<Rightarrow> 'b::t1_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2227 | assumes "connected s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2228 | and "continuous_on s f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2229 |     and "open {x \<in> s. f x = a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2230 | and "\<exists>x \<in> s. f x = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2231 | shows "\<forall>x \<in> s. f x = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2232 | using continuous_levelset_openin[OF assms(1,2), of a, unfolded openin_open] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2233 | using assms (3,4) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2234 | by fast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2235 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2236 | text \<open>Some arithmetical combinations (more to prove).\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2237 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2238 | lemma open_scaling[intro]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2239 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2240 | assumes "c \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2241 | and "open s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2242 | shows "open((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2243 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2244 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2245 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2246 | assume "x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2247 | then obtain e where "e>0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2248 | and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2249 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2250 | have "e * \<bar>c\<bar> > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2251 | using assms(1)[unfolded zero_less_abs_iff[symmetric]] \<open>e>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2252 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2253 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2254 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2255 | assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2256 | then have "norm ((1 / c) *\<^sub>R y - x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2257 | unfolding dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2258 | using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2259 | assms(1)[unfolded zero_less_abs_iff[symmetric]] by (simp del:zero_less_abs_iff) | 
| 67399 | 2260 | then have "y \<in> ( *\<^sub>R) c ` s" | 
| 2261 | using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "( *\<^sub>R) c"] | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2262 | using e[THEN spec[where x="(1 / c) *\<^sub>R y"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2263 | using assms(1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2264 | unfolding dist_norm scaleR_scaleR | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2265 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2266 | } | 
| 67399 | 2267 | ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> ( *\<^sub>R) c ` s" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2268 | apply (rule_tac x="e * \<bar>c\<bar>" in exI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2269 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2270 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2271 | then show ?thesis unfolding open_dist by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2272 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2273 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2274 | lemma minus_image_eq_vimage: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2275 | fixes A :: "'a::ab_group_add set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2276 | shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2277 | by (auto intro!: image_eqI [where f="\<lambda>x. - x"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2278 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2279 | lemma open_negations: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2280 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2281 | shows "open S \<Longrightarrow> open ((\<lambda>x. - x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2282 | using open_scaling [of "- 1" S] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2283 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2284 | lemma open_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2285 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2286 | assumes "open S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2287 | shows "open((\<lambda>x. a + x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2288 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2289 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2290 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2291 | have "continuous (at x) (\<lambda>x. x - a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2292 | by (intro continuous_diff continuous_ident continuous_const) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2293 | } | 
| 67399 | 2294 |   moreover have "{x. x - a \<in> S} = (+) a ` S"
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2295 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2296 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2297 | by (metis assms continuous_open_vimage vimage_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2298 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2299 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2300 | lemma open_neg_translation: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2301 | fixes s :: "'a::real_normed_vector set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2302 | assumes "open s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2303 | shows "open((\<lambda>x. a - x) ` s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2304 | using open_translation[OF open_negations[OF assms], of a] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2305 | by (auto simp: image_image) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2306 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2307 | lemma open_affinity: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2308 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2309 | assumes "open S" "c \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2310 | shows "open ((\<lambda>x. a + c *\<^sub>R x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2311 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2312 | have *: "(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2313 | unfolding o_def .. | 
| 67399 | 2314 | have "(+) a ` ( *\<^sub>R) c ` S = ((+) a \<circ> ( *\<^sub>R) c) ` S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2315 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2316 | then show ?thesis | 
| 67399 | 2317 | using assms open_translation[of "( *\<^sub>R) c ` S" a] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2318 | unfolding * | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2319 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2320 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2321 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2322 | lemma interior_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2323 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2324 | shows "interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (interior S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2325 | proof (rule set_eqI, rule) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2326 | fix x | 
| 67399 | 2327 | assume "x \<in> interior ((+) a ` S)" | 
| 2328 | then obtain e where "e > 0" and e: "ball x e \<subseteq> (+) a ` S" | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2329 | unfolding mem_interior by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2330 | then have "ball (x - a) e \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2331 | unfolding subset_eq Ball_def mem_ball dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2332 | by (auto simp: diff_diff_eq) | 
| 67399 | 2333 | then show "x \<in> (+) a ` interior S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2334 | unfolding image_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2335 | apply (rule_tac x="x - a" in bexI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2336 | unfolding mem_interior | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2337 | using \<open>e > 0\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2338 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2339 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2340 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2341 | fix x | 
| 67399 | 2342 | assume "x \<in> (+) a ` interior S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2343 | then obtain y e where "e > 0" and e: "ball y e \<subseteq> S" and y: "x = a + y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2344 | unfolding image_iff Bex_def mem_interior by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2345 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2346 | fix z | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2347 | have *: "a + y - z = y + a - z" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2348 | assume "z \<in> ball x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2349 | then have "z - a \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2350 | using e[unfolded subset_eq, THEN bspec[where x="z - a"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2351 | unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2352 | by auto | 
| 67399 | 2353 | then have "z \<in> (+) a ` S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2354 | unfolding image_iff by (auto intro!: bexI[where x="z - a"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2355 | } | 
| 67399 | 2356 | then have "ball x e \<subseteq> (+) a ` S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2357 | unfolding subset_eq by auto | 
| 67399 | 2358 | then show "x \<in> interior ((+) a ` S)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2359 | unfolding mem_interior using \<open>e > 0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2360 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2361 | |
| 67968 | 2362 | subsection \<open>Continuity implies uniform continuity on a compact domain\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2363 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2364 | text\<open>From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2365 | J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2366 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2367 | lemma Heine_Borel_lemma: | 
| 67237 | 2368 | assumes "compact S" and Ssub: "S \<subseteq> \<Union>\<G>" and opn: "\<And>G. G \<in> \<G> \<Longrightarrow> open G" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2369 | obtains e where "0 < e" "\<And>x. x \<in> S \<Longrightarrow> \<exists>G \<in> \<G>. ball x e \<subseteq> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2370 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2371 | have False if neg: "\<And>e. 0 < e \<Longrightarrow> \<exists>x \<in> S. \<forall>G \<in> \<G>. \<not> ball x e \<subseteq> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2372 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2373 | have "\<exists>x \<in> S. \<forall>G \<in> \<G>. \<not> ball x (1 / Suc n) \<subseteq> G" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2374 | using neg by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2375 | then obtain f where "\<And>n. f n \<in> S" and fG: "\<And>G n. G \<in> \<G> \<Longrightarrow> \<not> ball (f n) (1 / Suc n) \<subseteq> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2376 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2377 | then obtain l r where "l \<in> S" "strict_mono r" and to_l: "(f \<circ> r) \<longlonglongrightarrow> l" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2378 | using \<open>compact S\<close> compact_def that by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2379 | then obtain G where "l \<in> G" "G \<in> \<G>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2380 | using Ssub by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2381 | then obtain e where "0 < e" and e: "\<And>z. dist z l < e \<Longrightarrow> z \<in> G" | 
| 67237 | 2382 | using opn open_dist by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2383 | obtain N1 where N1: "\<And>n. n \<ge> N1 \<Longrightarrow> dist (f (r n)) l < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2384 | using to_l apply (simp add: lim_sequentially) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2385 | using \<open>0 < e\<close> half_gt_zero that by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2386 | obtain N2 where N2: "of_nat N2 > 2/e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2387 | using reals_Archimedean2 by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2388 | obtain x where "x \<in> ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \<notin> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2389 | using fG [OF \<open>G \<in> \<G>\<close>, of "r (max N1 N2)"] by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2390 | then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2391 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2392 | also have "... \<le> 1 / real (Suc (max N1 N2))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2393 | apply (simp add: divide_simps del: max.bounded_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2394 | using \<open>strict_mono r\<close> seq_suble by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2395 | also have "... \<le> 1 / real (Suc N2)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2396 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2397 | also have "... < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2398 | using N2 \<open>0 < e\<close> by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2399 | finally have "dist (f (r (max N1 N2))) x < e / 2" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2400 | moreover have "dist (f (r (max N1 N2))) l < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2401 | using N1 max.cobounded1 by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2402 | ultimately have "dist x l < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2403 | using dist_triangle_half_r by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2404 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2405 | using e \<open>x \<notin> G\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2406 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2407 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2408 | by (meson that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2409 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2410 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2411 | lemma compact_uniformly_equicontinuous: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2412 | assumes "compact S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2413 | and cont: "\<And>x e. \<lbrakk>x \<in> S; 0 < e\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2414 | \<Longrightarrow> \<exists>d. 0 < d \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2415 | (\<forall>f \<in> \<F>. \<forall>x' \<in> S. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2416 | and "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2417 | obtains d where "0 < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2418 | "\<And>f x x'. \<lbrakk>f \<in> \<F>; x \<in> S; x' \<in> S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2419 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2420 | obtain d where d_pos: "\<And>x e. \<lbrakk>x \<in> S; 0 < e\<rbrakk> \<Longrightarrow> 0 < d x e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2421 | and d_dist : "\<And>x x' e f. \<lbrakk>dist x' x < d x e; x \<in> S; x' \<in> S; 0 < e; f \<in> \<F>\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2422 | using cont by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2423 | let ?\<G> = "((\<lambda>x. ball x (d x (e / 2))) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2424 | have Ssub: "S \<subseteq> \<Union> ?\<G>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2425 | by clarsimp (metis d_pos \<open>0 < e\<close> dist_self half_gt_zero_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2426 | then obtain k where "0 < k" and k: "\<And>x. x \<in> S \<Longrightarrow> \<exists>G \<in> ?\<G>. ball x k \<subseteq> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2427 | by (rule Heine_Borel_lemma [OF \<open>compact S\<close>]) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2428 | moreover have "dist (f v) (f u) < e" if "f \<in> \<F>" "u \<in> S" "v \<in> S" "dist v u < k" for f u v | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2429 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2430 | obtain G where "G \<in> ?\<G>" "u \<in> G" "v \<in> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2431 | using k that | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2432 | by (metis \<open>dist v u < k\<close> \<open>u \<in> S\<close> \<open>0 < k\<close> centre_in_ball subsetD dist_commute mem_ball) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2433 | then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2434 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2435 | with that d_dist have "dist (f w) (f v) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2436 | by (metis \<open>0 < e\<close> dist_commute half_gt_zero) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2437 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2438 | have "dist (f w) (f u) < e/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2439 | using that d_dist w by (metis \<open>0 < e\<close> dist_commute divide_pos_pos zero_less_numeral) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2440 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2441 | using dist_triangle_half_r by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2442 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2443 | ultimately show ?thesis using that by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2444 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2445 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2446 | corollary compact_uniformly_continuous: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2447 | fixes f :: "'a :: metric_space \<Rightarrow> 'b :: metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2448 | assumes f: "continuous_on S f" and S: "compact S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2449 | shows "uniformly_continuous_on S f" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2450 | using f | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2451 | unfolding continuous_on_iff uniformly_continuous_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2452 |     by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"])
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2453 | |
| 67962 | 2454 | subsection%unimportant \<open>Topological stuff about the set of Reals\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2455 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2456 | lemma open_real: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2457 | fixes s :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2458 | shows "open s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. \<bar>x' - x\<bar> < e --> x' \<in> s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2459 | unfolding open_dist dist_norm by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2460 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2461 | lemma islimpt_approachable_real: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2462 | fixes s :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2463 | shows "x islimpt s \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> s. x' \<noteq> x \<and> \<bar>x' - x\<bar> < e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2464 | unfolding islimpt_approachable dist_norm by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2465 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2466 | lemma closed_real: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2467 | fixes s :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2468 | shows "closed s \<longleftrightarrow> (\<forall>x. (\<forall>e>0. \<exists>x' \<in> s. x' \<noteq> x \<and> \<bar>x' - x\<bar> < e) \<longrightarrow> x \<in> s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2469 | unfolding closed_limpt islimpt_approachable dist_norm by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2470 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2471 | lemma continuous_at_real_range: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2472 | fixes f :: "'a::real_normed_vector \<Rightarrow> real" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2473 | shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'. norm(x' - x) < d --> \<bar>f x' - f x\<bar> < e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2474 | unfolding continuous_at | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2475 | unfolding Lim_at | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2476 | unfolding dist_norm | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2477 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2478 | apply (erule_tac x=e in allE, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2479 | apply (rule_tac x=d in exI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2480 | apply (erule_tac x=x' in allE, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2481 | apply (erule_tac x=e in allE, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2482 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2483 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2484 | lemma continuous_on_real_range: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2485 | fixes f :: "'a::real_normed_vector \<Rightarrow> real" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2486 | shows "continuous_on s f \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2487 | (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2488 | unfolding continuous_on_iff dist_norm by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2489 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2490 | |
| 67962 | 2491 | subsection%unimportant \<open>Cartesian products\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2492 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2493 | lemma bounded_Times: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2494 | assumes "bounded s" "bounded t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2495 | shows "bounded (s \<times> t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2496 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2497 | obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2498 | using assms [unfolded bounded_def] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2499 | then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<^sup>2 + b\<^sup>2)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2500 | by (auto simp: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2501 | then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2502 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2503 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2504 | lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2505 | by (induct x) simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2506 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2507 | lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2508 | unfolding seq_compact_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2509 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2510 | apply (drule_tac x="fst \<circ> f" in spec) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2511 | apply (drule mp, simp add: mem_Times_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2512 | apply (clarify, rename_tac l1 r1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2513 | apply (drule_tac x="snd \<circ> f \<circ> r1" in spec) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2514 | apply (drule mp, simp add: mem_Times_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2515 | apply (clarify, rename_tac l2 r2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2516 | apply (rule_tac x="(l1, l2)" in rev_bexI, simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2517 | apply (rule_tac x="r1 \<circ> r2" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2518 | apply (rule conjI, simp add: strict_mono_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2519 | apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2520 | apply (drule (1) tendsto_Pair) back | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2521 | apply (simp add: o_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2522 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2523 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2524 | lemma compact_Times: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2525 | assumes "compact s" "compact t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2526 | shows "compact (s \<times> t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2527 | proof (rule compactI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2528 | fix C | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2529 | assume C: "\<forall>t\<in>C. open t" "s \<times> t \<subseteq> \<Union>C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2530 | have "\<forall>x\<in>s. \<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2531 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2532 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2533 | assume "x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2534 | have "\<forall>y\<in>t. \<exists>a b c. c \<in> C \<and> open a \<and> open b \<and> x \<in> a \<and> y \<in> b \<and> a \<times> b \<subseteq> c" (is "\<forall>y\<in>t. ?P y") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2535 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2536 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2537 | assume "y \<in> t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2538 | with \<open>x \<in> s\<close> C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2539 | then show "?P y" by (auto elim!: open_prod_elim) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2540 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2541 | then obtain a b c where b: "\<And>y. y \<in> t \<Longrightarrow> open (b y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2542 | and c: "\<And>y. y \<in> t \<Longrightarrow> c y \<in> C \<and> open (a y) \<and> open (b y) \<and> x \<in> a y \<and> y \<in> b y \<and> a y \<times> b y \<subseteq> c y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2543 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2544 | then have "\<forall>y\<in>t. open (b y)" "t \<subseteq> (\<Union>y\<in>t. b y)" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2545 | with compactE_image[OF \<open>compact t\<close>] obtain D where D: "D \<subseteq> t" "finite D" "t \<subseteq> (\<Union>y\<in>D. b y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2546 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2547 | moreover from D c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2548 | by (fastforce simp: subset_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2549 | ultimately show "\<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2550 | using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>(a`D)"] conjI) (auto intro!: open_INT) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2551 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2552 | then obtain a d where a: "\<And>x. x\<in>s \<Longrightarrow> open (a x)" "s \<subseteq> (\<Union>x\<in>s. a x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2553 | and d: "\<And>x. x \<in> s \<Longrightarrow> d x \<subseteq> C \<and> finite (d x) \<and> a x \<times> t \<subseteq> \<Union>d x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2554 | unfolding subset_eq UN_iff by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2555 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2556 | from compactE_image[OF \<open>compact s\<close> a] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2557 | obtain e where e: "e \<subseteq> s" "finite e" and s: "s \<subseteq> (\<Union>x\<in>e. a x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2558 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2559 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2560 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2561 | from s have "s \<times> t \<subseteq> (\<Union>x\<in>e. a x \<times> t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2562 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2563 | also have "\<dots> \<subseteq> (\<Union>x\<in>e. \<Union>d x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2564 | using d \<open>e \<subseteq> s\<close> by (intro UN_mono) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2565 | finally have "s \<times> t \<subseteq> (\<Union>x\<in>e. \<Union>d x)" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2566 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2567 | ultimately show "\<exists>C'\<subseteq>C. finite C' \<and> s \<times> t \<subseteq> \<Union>C'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2568 | by (intro exI[of _ "(\<Union>x\<in>e. d x)"]) (auto simp: subset_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2569 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2570 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2571 | text\<open>Hence some useful properties follow quite easily.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2572 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2573 | lemma compact_scaling: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2574 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2575 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2576 | shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2577 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2578 | let ?f = "\<lambda>x. scaleR c x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2579 | have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2580 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2581 | using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2582 | using linear_continuous_at[OF *] assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2583 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2584 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2585 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2586 | lemma compact_negations: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2587 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2588 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2589 | shows "compact ((\<lambda>x. - x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2590 | using compact_scaling [OF assms, of "- 1"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2591 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2592 | lemma compact_sums: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2593 | fixes s t :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2594 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2595 | and "compact t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2596 |   shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2597 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2598 |   have *: "{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2599 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2600 | unfolding image_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2601 | apply (rule_tac x="(xa, y)" in bexI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2602 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2603 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2604 | have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2605 | unfolding continuous_on by (rule ballI) (intro tendsto_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2606 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2607 | unfolding * using compact_continuous_image compact_Times [OF assms] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2608 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2609 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2610 | lemma compact_differences: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2611 | fixes s t :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2612 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2613 | and "compact t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2614 |   shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2615 | proof- | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2616 |   have "{x - y | x y. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2617 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2618 | apply (rule_tac x= xa in exI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2619 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2620 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2621 | using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2622 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2623 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2624 | lemma compact_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2625 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2626 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2627 | shows "compact ((\<lambda>x. a + x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2628 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2629 |   have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2630 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2631 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2632 | using compact_sums[OF assms compact_sing[of a]] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2633 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2634 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2635 | lemma compact_affinity: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2636 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2637 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2638 | shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2639 | proof - | 
| 67399 | 2640 | have "(+) a ` ( *\<^sub>R) c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2641 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2642 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2643 | using compact_translation[OF compact_scaling[OF assms], of a c] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2644 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2645 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2646 | text \<open>Hence we get the following.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2647 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2648 | lemma compact_sup_maxdistance: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2649 | fixes s :: "'a::metric_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2650 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2651 |     and "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2652 | shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2653 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2654 | have "compact (s \<times> s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2655 | using \<open>compact s\<close> by (intro compact_Times) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2656 |   moreover have "s \<times> s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2657 |     using \<open>s \<noteq> {}\<close> by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2658 | moreover have "continuous_on (s \<times> s) (\<lambda>x. dist (fst x) (snd x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2659 | by (intro continuous_at_imp_continuous_on ballI continuous_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2660 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2661 | using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2662 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2663 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2664 | |
| 67968 | 2665 | subsection \<open>The diameter of a set\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2666 | |
| 67962 | 2667 | definition%important diameter :: "'a::metric_space set \<Rightarrow> real" where | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2668 |   "diameter S = (if S = {} then 0 else SUP (x,y):S\<times>S. dist x y)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2669 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2670 | lemma diameter_empty [simp]: "diameter{} = 0"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2671 | by (auto simp: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2672 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2673 | lemma diameter_singleton [simp]: "diameter{x} = 0"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2674 | by (auto simp: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2675 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2676 | lemma diameter_le: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2677 |   assumes "S \<noteq> {} \<or> 0 \<le> d"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2678 | and no: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> norm(x - y) \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2679 | shows "diameter S \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2680 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2681 | by (auto simp: dist_norm diameter_def intro: cSUP_least) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2682 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2683 | lemma diameter_bounded_bound: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2684 | fixes s :: "'a :: metric_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2685 | assumes s: "bounded s" "x \<in> s" "y \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2686 | shows "dist x y \<le> diameter s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2687 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2688 | from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2689 | unfolding bounded_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2690 | have "bdd_above (case_prod dist ` (s\<times>s))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2691 | proof (intro bdd_aboveI, safe) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2692 | fix a b | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2693 | assume "a \<in> s" "b \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2694 | with z[of a] z[of b] dist_triangle[of a b z] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2695 | show "dist a b \<le> 2 * d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2696 | by (simp add: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2697 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2698 | moreover have "(x,y) \<in> s\<times>s" using s by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2699 | ultimately have "dist x y \<le> (SUP (x,y):s\<times>s. dist x y)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2700 | by (rule cSUP_upper2) simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2701 | with \<open>x \<in> s\<close> show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2702 | by (auto simp: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2703 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2704 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2705 | lemma diameter_lower_bounded: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2706 | fixes s :: "'a :: metric_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2707 | assumes s: "bounded s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2708 | and d: "0 < d" "d < diameter s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2709 | shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2710 | proof (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2711 | assume contr: "\<not> ?thesis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2712 |   moreover have "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2713 | using d by (auto simp: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2714 | ultimately have "diameter s \<le> d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2715 | by (auto simp: not_less diameter_def intro!: cSUP_least) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2716 | with \<open>d < diameter s\<close> show False by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2717 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2718 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2719 | lemma diameter_bounded: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2720 | assumes "bounded s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2721 | shows "\<forall>x\<in>s. \<forall>y\<in>s. dist x y \<le> diameter s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2722 | and "\<forall>d>0. d < diameter s \<longrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. dist x y > d)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2723 | using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2724 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2725 | |
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2726 | lemma bounded_two_points: | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2727 | "bounded S \<longleftrightarrow> (\<exists>e. \<forall>x\<in>S. \<forall>y\<in>S. dist x y \<le> e)" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2728 | apply (rule iffI) | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2729 | subgoal using diameter_bounded(1) by auto | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2730 | subgoal using bounded_any_center[of S] by meson | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2731 | done | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 2732 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2733 | lemma diameter_compact_attained: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2734 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2735 |     and "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2736 | shows "\<exists>x\<in>s. \<exists>y\<in>s. dist x y = diameter s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2737 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2738 | have b: "bounded s" using assms(1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2739 | by (rule compact_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2740 | then obtain x y where xys: "x\<in>s" "y\<in>s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2741 | and xy: "\<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2742 | using compact_sup_maxdistance[OF assms] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2743 | then have "diameter s \<le> dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2744 | unfolding diameter_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2745 | apply clarsimp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2746 | apply (rule cSUP_least, fast+) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2747 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2748 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2749 | by (metis b diameter_bounded_bound order_antisym xys) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2750 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2751 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2752 | lemma diameter_ge_0: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2753 | assumes "bounded S" shows "0 \<le> diameter S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2754 | by (metis all_not_in_conv assms diameter_bounded_bound diameter_empty dist_self order_refl) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2755 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2756 | lemma diameter_subset: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2757 | assumes "S \<subseteq> T" "bounded T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2758 | shows "diameter S \<le> diameter T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2759 | proof (cases "S = {} \<or> T = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2760 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2761 | with assms show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2762 | by (force simp: diameter_ge_0) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2763 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2764 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2765 | then have "bdd_above ((\<lambda>x. case x of (x, xa) \<Rightarrow> dist x xa) ` (T \<times> T))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2766 | using \<open>bounded T\<close> diameter_bounded_bound by (force simp: bdd_above_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2767 | with False \<open>S \<subseteq> T\<close> show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2768 | apply (simp add: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2769 | apply (rule cSUP_subset_mono, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2770 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2771 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2772 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2773 | lemma diameter_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2774 | assumes "bounded S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2775 | shows "diameter(closure S) = diameter S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2776 | proof (rule order_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2777 | have "False" if "diameter S < diameter (closure S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2778 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2779 | define d where "d = diameter(closure S) - diameter(S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2780 | have "d > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2781 | using that by (simp add: d_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2782 | then have "diameter(closure(S)) - d / 2 < diameter(closure(S))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2783 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2784 | have dd: "diameter (closure S) - d / 2 = (diameter(closure(S)) + diameter(S)) / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2785 | by (simp add: d_def divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2786 | have bocl: "bounded (closure S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2787 | using assms by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2788 | moreover have "0 \<le> diameter S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2789 | using assms diameter_ge_0 by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2790 | ultimately obtain x y where "x \<in> closure S" "y \<in> closure S" and xy: "diameter(closure(S)) - d / 2 < dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2791 | using diameter_bounded(2) [OF bocl, rule_format, of "diameter(closure(S)) - d / 2"] \<open>d > 0\<close> d_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2792 | then obtain x' y' where x'y': "x' \<in> S" "dist x' x < d/4" "y' \<in> S" "dist y' y < d/4" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2793 | using closure_approachable | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2794 | by (metis \<open>0 < d\<close> zero_less_divide_iff zero_less_numeral) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2795 | then have "dist x' y' \<le> diameter S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2796 | using assms diameter_bounded_bound by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2797 | with x'y' have "dist x y \<le> d / 4 + diameter S + d / 4" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2798 | by (meson add_mono_thms_linordered_semiring(1) dist_triangle dist_triangle3 less_eq_real_def order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2799 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2800 | using xy d_def by linarith | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2801 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2802 | then show "diameter (closure S) \<le> diameter S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2803 | by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2804 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2805 | show "diameter S \<le> diameter (closure S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2806 | by (simp add: assms bounded_closure closure_subset diameter_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2807 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2808 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2809 | lemma diameter_cball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2810 | fixes a :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2811 | shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2812 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2813 | have "diameter(cball a r) = 2*r" if "r \<ge> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2814 | proof (rule order_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2815 | show "diameter (cball a r) \<le> 2*r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2816 | proof (rule diameter_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2817 | fix x y assume "x \<in> cball a r" "y \<in> cball a r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2818 | then have "norm (x - a) \<le> r" "norm (a - y) \<le> r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2819 | by (auto simp: dist_norm norm_minus_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2820 | then have "norm (x - y) \<le> r+r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2821 | using norm_diff_triangle_le by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2822 | then show "norm (x - y) \<le> 2*r" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2823 | qed (simp add: that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2824 | have "2*r = dist (a + r *\<^sub>R (SOME i. i \<in> Basis)) (a - r *\<^sub>R (SOME i. i \<in> Basis))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2825 | apply (simp add: dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2826 | by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2827 | also have "... \<le> diameter (cball a r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2828 | apply (rule diameter_bounded_bound) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2829 | using that by (auto simp: dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2830 | finally show "2*r \<le> diameter (cball a r)" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2831 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2832 | then show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2833 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2834 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2835 | lemma diameter_ball [simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2836 | fixes a :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2837 | shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2838 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2839 | have "diameter(ball a r) = 2*r" if "r > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2840 | by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2841 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2842 | by (simp add: diameter_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2843 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2844 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2845 | lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2846 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2847 |   have "{a .. b} = cball ((a+b)/2) ((b-a)/2)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2848 | by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2849 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2850 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2851 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2852 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2853 | lemma diameter_open_interval [simp]: "diameter {a<..<b} = (if b < a then 0 else b-a)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2854 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2855 |   have "{a <..< b} = ball ((a+b)/2) ((b-a)/2)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2856 | by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2857 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2858 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2859 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2860 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2861 | proposition Lebesgue_number_lemma: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2862 |   assumes "compact S" "\<C> \<noteq> {}" "S \<subseteq> \<Union>\<C>" and ope: "\<And>B. B \<in> \<C> \<Longrightarrow> open B"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2863 | obtains \<delta> where "0 < \<delta>" "\<And>T. \<lbrakk>T \<subseteq> S; diameter T < \<delta>\<rbrakk> \<Longrightarrow> \<exists>B \<in> \<C>. T \<subseteq> B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2864 | proof (cases "S = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2865 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2866 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2867 |     by (metis \<open>\<C> \<noteq> {}\<close> zero_less_one empty_subsetI equals0I subset_trans that)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2868 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2869 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2870 |   { fix x assume "x \<in> S"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2871 | then obtain C where C: "x \<in> C" "C \<in> \<C>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2872 | using \<open>S \<subseteq> \<Union>\<C>\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2873 | then obtain r where r: "r>0" "ball x (2*r) \<subseteq> C" | 
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
68302diff
changeset | 2874 | by (metis mult.commute mult_2_right not_le ope openE field_sum_of_halves zero_le_numeral zero_less_mult_iff) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2875 | then have "\<exists>r C. r > 0 \<and> ball x (2*r) \<subseteq> C \<and> C \<in> \<C>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2876 | using C by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2877 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2878 | then obtain r where r: "\<And>x. x \<in> S \<Longrightarrow> r x > 0 \<and> (\<exists>C \<in> \<C>. ball x (2*r x) \<subseteq> C)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2879 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2880 | then have "S \<subseteq> (\<Union>x \<in> S. ball x (r x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2881 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2882 | then obtain \<T> where "finite \<T>" "S \<subseteq> \<Union>\<T>" and \<T>: "\<T> \<subseteq> (\<lambda>x. ball x (r x)) ` S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2883 | by (rule compactE [OF \<open>compact S\<close>]) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2884 | then obtain S0 where "S0 \<subseteq> S" "finite S0" and S0: "\<T> = (\<lambda>x. ball x (r x)) ` S0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2885 | by (meson finite_subset_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2886 |   then have "S0 \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2887 | using False \<open>S \<subseteq> \<Union>\<T>\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2888 | define \<delta> where "\<delta> = Inf (r ` S0)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2889 | have "\<delta> > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2890 |     using \<open>finite S0\<close> \<open>S0 \<subseteq> S\<close> \<open>S0 \<noteq> {}\<close> r by (auto simp: \<delta>_def finite_less_Inf_iff)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2891 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2892 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2893 | show "0 < \<delta>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2894 | by (simp add: \<open>0 < \<delta>\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2895 | show "\<exists>B \<in> \<C>. T \<subseteq> B" if "T \<subseteq> S" and dia: "diameter T < \<delta>" for T | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2896 |     proof (cases "T = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2897 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2898 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2899 |         using \<open>\<C> \<noteq> {}\<close> by blast
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2900 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2901 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2902 | then obtain y where "y \<in> T" by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2903 | then have "y \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2904 | using \<open>T \<subseteq> S\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2905 | then obtain x where "x \<in> S0" and x: "y \<in> ball x (r x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2906 | using \<open>S \<subseteq> \<Union>\<T>\<close> S0 that by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2907 | have "ball y \<delta> \<subseteq> ball y (r x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2908 |         by (metis \<delta>_def \<open>S0 \<noteq> {}\<close> \<open>finite S0\<close> \<open>x \<in> S0\<close> empty_is_image finite_imageI finite_less_Inf_iff imageI less_irrefl not_le subset_ball)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2909 | also have "... \<subseteq> ball x (2*r x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2910 | by clarsimp (metis dist_commute dist_triangle_less_add mem_ball mult_2 x) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2911 | finally obtain C where "C \<in> \<C>" "ball y \<delta> \<subseteq> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2912 | by (meson r \<open>S0 \<subseteq> S\<close> \<open>x \<in> S0\<close> dual_order.trans subsetCE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2913 | have "bounded T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2914 | using \<open>compact S\<close> bounded_subset compact_imp_bounded \<open>T \<subseteq> S\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2915 | then have "T \<subseteq> ball y \<delta>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2916 | using \<open>y \<in> T\<close> dia diameter_bounded_bound by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2917 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2918 | apply (rule_tac x=C in bexI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2919 | using \<open>ball y \<delta> \<subseteq> C\<close> \<open>C \<in> \<C>\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2920 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2921 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2922 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2923 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2924 | lemma diameter_cbox: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2925 | fixes a b::"'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2926 | shows "(\<forall>i \<in> Basis. a \<bullet> i \<le> b \<bullet> i) \<Longrightarrow> diameter (cbox a b) = dist a b" | 
| 67155 | 2927 | by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2928 | simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2929 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2930 | subsection \<open>Separation between points and sets\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2931 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2932 | proposition separate_point_closed: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2933 | fixes s :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2934 | assumes "closed s" and "a \<notin> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2935 | shows "\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2936 | proof (cases "s = {}")
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2937 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2938 | then show ?thesis by(auto intro!: exI[where x=1]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2939 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2940 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2941 | from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2942 |     using \<open>s \<noteq> {}\<close> by (blast intro: distance_attains_inf [of s a])
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2943 | with \<open>x\<in>s\<close> show ?thesis using dist_pos_lt[of a x] and\<open>a \<notin> s\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2944 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2945 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2946 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2947 | proposition separate_compact_closed: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2948 | fixes s t :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2949 | assumes "compact s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2950 |     and t: "closed t" "s \<inter> t = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2951 | shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2952 | proof cases | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2953 |   assume "s \<noteq> {} \<and> t \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2954 |   then have "s \<noteq> {}" "t \<noteq> {}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2955 | let ?inf = "\<lambda>x. infdist x t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2956 | have "continuous_on s ?inf" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2957 | by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_ident) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2958 | then obtain x where x: "x \<in> s" "\<forall>y\<in>s. ?inf x \<le> ?inf y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2959 |     using continuous_attains_inf[OF \<open>compact s\<close> \<open>s \<noteq> {}\<close>] by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2960 | then have "0 < ?inf x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2961 |     using t \<open>t \<noteq> {}\<close> in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2962 | moreover have "\<forall>x'\<in>s. \<forall>y\<in>t. ?inf x \<le> dist x' y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2963 | using x by (auto intro: order_trans infdist_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2964 | ultimately show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2965 | qed (auto intro!: exI[of _ 1]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2966 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2967 | proposition separate_closed_compact: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2968 | fixes s t :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2969 | assumes "closed s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2970 | and "compact t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2971 |     and "s \<inter> t = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2972 | shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2973 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2974 |   have *: "t \<inter> s = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2975 | using assms(3) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2976 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2977 | using separate_compact_closed[OF assms(2,1) *] by (force simp: dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2978 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2979 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2980 | proposition compact_in_open_separated: | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2981 | fixes A::"'a::heine_borel set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2982 |   assumes "A \<noteq> {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2983 | assumes "compact A" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2984 | assumes "open B" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2985 | assumes "A \<subseteq> B" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2986 |   obtains e where "e > 0" "{x. infdist x A \<le> e} \<subseteq> B"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 2987 | proof atomize_elim | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2988 |   have "closed (- B)" "compact A" "- B \<inter> A = {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2989 | using assms by (auto simp: open_Diff compact_eq_bounded_closed) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2990 | from separate_closed_compact[OF this] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2991 | obtain d'::real where d': "d'>0" "\<And>x y. x \<notin> B \<Longrightarrow> y \<in> A \<Longrightarrow> d' \<le> dist x y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2992 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2993 | define d where "d = d' / 2" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2994 | hence "d>0" "d < d'" using d' by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2995 | with d' have d: "\<And>x y. x \<notin> B \<Longrightarrow> y \<in> A \<Longrightarrow> d < dist x y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2996 | by force | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2997 |   show "\<exists>e>0. {x. infdist x A \<le> e} \<subseteq> B"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2998 | proof (rule ccontr) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 2999 |     assume "\<nexists>e. 0 < e \<and> {x. infdist x A \<le> e} \<subseteq> B"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3000 | with \<open>d > 0\<close> obtain x where x: "infdist x A \<le> d" "x \<notin> B" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3001 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3002 |     from assms have "closed A" "A \<noteq> {}" by (auto simp: compact_eq_bounded_closed)
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3003 | from infdist_attains_inf[OF this] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3004 | obtain y where y: "y \<in> A" "infdist x A = dist x y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3005 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3006 | have "dist x y \<le> d" using x y by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3007 | also have "\<dots> < dist x y" using y d x by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3008 | finally show False by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3009 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3010 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 3011 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3012 | |
| 67968 | 3013 | subsection%unimportant \<open>Compact sets and the closure operation\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3014 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3015 | lemma closed_scaling: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3016 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3017 | assumes "closed S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3018 | shows "closed ((\<lambda>x. c *\<^sub>R x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3019 | proof (cases "c = 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3020 | case True then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3021 | by (auto simp: image_constant_conv) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3022 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3023 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3024 | from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3025 | by (simp add: continuous_closed_vimage) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3026 | also have "(\<lambda>x. inverse c *\<^sub>R x) -` S = (\<lambda>x. c *\<^sub>R x) ` S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3027 | using \<open>c \<noteq> 0\<close> by (auto elim: image_eqI [rotated]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3028 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3029 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3030 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3031 | lemma closed_negations: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3032 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3033 | assumes "closed S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3034 | shows "closed ((\<lambda>x. -x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3035 | using closed_scaling[OF assms, of "- 1"] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3036 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3037 | lemma compact_closed_sums: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3038 | fixes S :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3039 | assumes "compact S" and "closed T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3040 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3041 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3042 |   let ?S = "{x + y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3043 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3044 | fix x l | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3045 | assume as: "\<forall>n. x n \<in> ?S" "(x \<longlongrightarrow> l) sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3046 | from as(1) obtain f where f: "\<forall>n. x n = fst (f n) + snd (f n)" "\<forall>n. fst (f n) \<in> S" "\<forall>n. snd (f n) \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3047 | using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> S \<and> snd y \<in> T"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3048 | obtain l' r where "l'\<in>S" and r: "strict_mono r" and lr: "(((\<lambda>n. fst (f n)) \<circ> r) \<longlongrightarrow> l') sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3049 | using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3050 | have "((\<lambda>n. snd (f (r n))) \<longlongrightarrow> l - l') sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3051 | using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3052 | unfolding o_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3053 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3054 | then have "l - l' \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3055 | using assms(2)[unfolded closed_sequential_limits, | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3056 | THEN spec[where x="\<lambda> n. snd (f (r n))"], | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3057 | THEN spec[where x="l - l'"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3058 | using f(3) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3059 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3060 | then have "l \<in> ?S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3061 | using \<open>l' \<in> S\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3062 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3063 | apply (rule_tac x=l' in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3064 | apply (rule_tac x="l - l'" in exI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3065 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3066 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3067 |   moreover have "?S = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3068 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3069 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3070 | unfolding closed_sequential_limits | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3071 | by (metis (no_types, lifting)) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3072 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3073 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3074 | lemma closed_compact_sums: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3075 | fixes S T :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3076 | assumes "closed S" "compact T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3077 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3078 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3079 |   have "(\<Union>x\<in> T. \<Union>y \<in> S. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3080 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3081 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3082 | using compact_closed_sums[OF assms(2,1)] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3083 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3084 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3085 | lemma compact_closed_differences: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3086 | fixes S T :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3087 | assumes "compact S" "closed T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3088 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3089 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3090 |   have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3091 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3092 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3093 | using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3094 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3095 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3096 | lemma closed_compact_differences: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3097 | fixes S T :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3098 | assumes "closed S" "compact T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3099 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3100 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3101 |   have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = {x - y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3102 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3103 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3104 | using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3105 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3106 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3107 | lemma closed_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3108 | fixes a :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3109 | assumes "closed S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3110 | shows "closed ((\<lambda>x. a + x) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3111 | proof - | 
| 67399 | 3112 |   have "(\<Union>x\<in> {a}. \<Union>y \<in> S. {x + y}) = ((+) a ` S)" by auto
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3113 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3114 | using compact_closed_sums[OF compact_sing[of a] assms] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3115 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3116 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3117 | lemma closure_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3118 | fixes a :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3119 | shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3120 | proof - | 
| 67399 | 3121 | have *: "(+) a ` (- s) = - (+) a ` s" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3122 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3123 | unfolding image_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3124 | apply (rule_tac x="x - a" in bexI, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3125 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3126 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3127 | unfolding closure_interior translation_Compl | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3128 | using interior_translation[of a "- s"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3129 | unfolding * | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3130 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3131 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3132 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3133 | lemma frontier_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3134 | fixes a :: "'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3135 | shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3136 | unfolding frontier_def translation_diff interior_translation closure_translation | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3137 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3138 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3139 | lemma sphere_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3140 | fixes a :: "'n::euclidean_space" | 
| 67399 | 3141 | shows "sphere (a+c) r = (+) a ` sphere c r" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3142 | apply safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3143 | apply (rule_tac x="x-a" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3144 | apply (auto simp: dist_norm algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3145 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3146 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3147 | lemma cball_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3148 | fixes a :: "'n::euclidean_space" | 
| 67399 | 3149 | shows "cball (a+c) r = (+) a ` cball c r" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3150 | apply safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3151 | apply (rule_tac x="x-a" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3152 | apply (auto simp: dist_norm algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3153 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3154 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3155 | lemma ball_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3156 | fixes a :: "'n::euclidean_space" | 
| 67399 | 3157 | shows "ball (a+c) r = (+) a ` ball c r" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3158 | apply safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3159 | apply (rule_tac x="x-a" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3160 | apply (auto simp: dist_norm algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3161 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3162 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3163 | |
| 67962 | 3164 | subsection%unimportant \<open>Closure of halfspaces and hyperplanes\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3165 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3166 | lemma continuous_on_closed_Collect_le: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3167 | fixes f g :: "'a::t2_space \<Rightarrow> real" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3168 | assumes f: "continuous_on s f" and g: "continuous_on s g" and s: "closed s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3169 |   shows "closed {x \<in> s. f x \<le> g x}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3170 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3171 |   have "closed ((\<lambda>x. g x - f x) -` {0..} \<inter> s)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3172 | using closed_real_atLeast continuous_on_diff [OF g f] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3173 | by (simp add: continuous_on_closed_vimage [OF s]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3174 |   also have "((\<lambda>x. g x - f x) -` {0..} \<inter> s) = {x\<in>s. f x \<le> g x}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3175 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3176 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3177 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3178 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3179 | lemma continuous_at_inner: "continuous (at x) (inner a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3180 | unfolding continuous_at by (intro tendsto_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3181 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3182 | lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3183 | by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3184 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3185 | lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3186 | by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3187 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3188 | lemma closed_hyperplane: "closed {x. inner a x = b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3189 | by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3190 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3191 | lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3192 | by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3193 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3194 | lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3195 | by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3196 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3197 | lemma closed_interval_left: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3198 | fixes b :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3199 |   shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3200 | by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3201 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3202 | lemma closed_interval_right: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3203 | fixes a :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3204 |   shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3205 | by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3206 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3207 | lemma continuous_le_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3208 | fixes a::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3209 | assumes f: "continuous_on (closure s) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3210 | and x: "x \<in> closure(s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3211 | and xlo: "\<And>x. x \<in> s ==> f(x) \<le> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3212 | shows "f(x) \<le> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3213 | using image_closure_subset [OF f] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3214 | using image_closure_subset [OF f] closed_halfspace_le [of "1::real" a] assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3215 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3216 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3217 | lemma continuous_ge_on_closure: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3218 | fixes a::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3219 | assumes f: "continuous_on (closure s) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3220 | and x: "x \<in> closure(s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3221 | and xlo: "\<And>x. x \<in> s ==> f(x) \<ge> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3222 | shows "f(x) \<ge> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3223 | using image_closure_subset [OF f] closed_halfspace_ge [of a "1::real"] assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3224 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3225 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3226 | lemma Lim_component_le: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3227 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3228 | assumes "(f \<longlongrightarrow> l) net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3229 | and "\<not> (trivial_limit net)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3230 | and "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3231 | shows "l\<bullet>i \<le> b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3232 | by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3233 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3234 | lemma Lim_component_ge: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3235 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3236 | assumes "(f \<longlongrightarrow> l) net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3237 | and "\<not> (trivial_limit net)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3238 | and "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3239 | shows "b \<le> l\<bullet>i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3240 | by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3241 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3242 | lemma Lim_component_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3243 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3244 | assumes net: "(f \<longlongrightarrow> l) net" "\<not> trivial_limit net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3245 | and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3246 | shows "l\<bullet>i = b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3247 | using ev[unfolded order_eq_iff eventually_conj_iff] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3248 | using Lim_component_ge[OF net, of b i] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3249 | using Lim_component_le[OF net, of i b] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3250 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3251 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3252 | text \<open>Limits relative to a union.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3253 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3254 | lemma eventually_within_Un: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3255 | "eventually P (at x within (s \<union> t)) \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3256 | eventually P (at x within s) \<and> eventually P (at x within t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3257 | unfolding eventually_at_filter | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3258 | by (auto elim!: eventually_rev_mp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3259 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3260 | lemma Lim_within_union: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3261 | "(f \<longlongrightarrow> l) (at x within (s \<union> t)) \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3262 | (f \<longlongrightarrow> l) (at x within s) \<and> (f \<longlongrightarrow> l) (at x within t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3263 | unfolding tendsto_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3264 | by (auto simp: eventually_within_Un) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3265 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3266 | lemma Lim_topological: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3267 | "(f \<longlongrightarrow> l) net \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3268 | trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3269 | unfolding tendsto_def trivial_limit_eq by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3270 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3271 | text \<open>Continuity relative to a union.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3272 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3273 | lemma continuous_on_Un_local: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3274 | "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3275 | continuous_on s f; continuous_on t f\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3276 | \<Longrightarrow> continuous_on (s \<union> t) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3277 | unfolding continuous_on closedin_limpt | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3278 | by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3279 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3280 | lemma continuous_on_cases_local: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3281 | "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3282 | continuous_on s f; continuous_on t g; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3283 | \<And>x. \<lbrakk>x \<in> s \<and> ~P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3284 | \<Longrightarrow> continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3285 | by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3286 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3287 | lemma continuous_on_cases_le: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3288 | fixes h :: "'a :: topological_space \<Rightarrow> real" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3289 |   assumes "continuous_on {t \<in> s. h t \<le> a} f"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3290 |       and "continuous_on {t \<in> s. a \<le> h t} g"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3291 | and h: "continuous_on s h" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3292 | and "\<And>t. \<lbrakk>t \<in> s; h t = a\<rbrakk> \<Longrightarrow> f t = g t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3293 | shows "continuous_on s (\<lambda>t. if h t \<le> a then f(t) else g(t))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3294 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3295 | have s: "s = (s \<inter> h -` atMost a) \<union> (s \<inter> h -` atLeast a)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3296 | by force | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3297 | have 1: "closedin (subtopology euclidean s) (s \<inter> h -` atMost a)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3298 | by (rule continuous_closedin_preimage [OF h closed_atMost]) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3299 | have 2: "closedin (subtopology euclidean s) (s \<inter> h -` atLeast a)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3300 | by (rule continuous_closedin_preimage [OF h closed_atLeast]) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3301 |   have eq: "s \<inter> h -` {..a} = {t \<in> s. h t \<le> a}" "s \<inter> h -` {a..} = {t \<in> s. a \<le> h t}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3302 | by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3303 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3304 | apply (rule continuous_on_subset [of s, OF _ order_refl]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3305 | apply (subst s) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3306 | apply (rule continuous_on_cases_local) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3307 | using 1 2 s assms apply (auto simp: eq) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3308 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3309 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3310 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3311 | lemma continuous_on_cases_1: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3312 | fixes s :: "real set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3313 |   assumes "continuous_on {t \<in> s. t \<le> a} f"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3314 |       and "continuous_on {t \<in> s. a \<le> t} g"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3315 | and "a \<in> s \<Longrightarrow> f a = g a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3316 | shows "continuous_on s (\<lambda>t. if t \<le> a then f(t) else g(t))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3317 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3318 | by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3319 | |
| 67968 | 3320 | subsubsection\<open>Some more convenient intermediate-value theorem formulations\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3321 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3322 | lemma connected_ivt_hyperplane: | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3323 | assumes "connected S" and xy: "x \<in> S" "y \<in> S" and b: "inner a x \<le> b" "b \<le> inner a y" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3324 | shows "\<exists>z \<in> S. inner a z = b" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3325 | proof (rule ccontr) | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3326 | assume as:"\<not> (\<exists>z\<in>S. inner a z = b)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3327 |   let ?A = "{x. inner a x < b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3328 |   let ?B = "{x. inner a x > b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3329 | have "open ?A" "open ?B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3330 | using open_halfspace_lt and open_halfspace_gt by auto | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3331 |   moreover have "?A \<inter> ?B = {}" by auto
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3332 | moreover have "S \<subseteq> ?A \<union> ?B" using as by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3333 | ultimately show False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3334 | using \<open>connected S\<close>[unfolded connected_def not_ex, | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3335 | THEN spec[where x="?A"], THEN spec[where x="?B"]] | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3336 | using xy b by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3337 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3338 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3339 | lemma connected_ivt_component: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3340 | fixes x::"'a::euclidean_space" | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3341 | shows "connected S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>S. z\<bullet>k = a)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 3342 | using connected_ivt_hyperplane[of S x y "k::'a" a] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3343 | by (auto simp: inner_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3344 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3345 | lemma image_affinity_cbox: fixes m::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3346 | fixes a b c :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3347 | shows "(\<lambda>x. m *\<^sub>R x + c) ` cbox a b = | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3348 |     (if cbox a b = {} then {}
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3349 | else (if 0 \<le> m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3350 | else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3351 | proof (cases "m = 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3352 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3353 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3354 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3355 | assume "\<forall>i\<in>Basis. x \<bullet> i \<le> c \<bullet> i" "\<forall>i\<in>Basis. c \<bullet> i \<le> x \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3356 | then have "x = c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3357 | by (simp add: dual_order.antisym euclidean_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3358 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3359 | moreover have "c \<in> cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3360 | unfolding True by (auto simp: cbox_sing) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3361 | ultimately show ?thesis using True by (auto simp: cbox_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3362 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3363 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3364 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3365 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3366 | assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3367 | then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3368 | by (auto simp: inner_distrib) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3369 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3370 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3371 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3372 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3373 | assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m < 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3374 | then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3375 | by (auto simp: mult_left_mono_neg inner_distrib) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3376 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3377 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3378 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3379 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3380 | assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i" and "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3381 | then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3382 | unfolding image_iff Bex_def mem_box | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3383 | apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3384 | apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3385 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3386 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3387 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3388 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3389 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3390 | assume "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" "m < 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3391 | then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3392 | unfolding image_iff Bex_def mem_box | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3393 | apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3394 | apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3395 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3396 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3397 | ultimately show ?thesis using False by (auto simp: cbox_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3398 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3399 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3400 | lemma image_smult_cbox:"(\<lambda>x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3401 |   (if cbox a b = {} then {} else if 0 \<le> m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3402 | using image_affinity_cbox[of m 0 a b] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3403 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3404 | lemma islimpt_greaterThanLessThan1: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3405 |   fixes a b::"'a::{linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3406 | assumes "a < b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3407 |   shows  "a islimpt {a<..<b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3408 | proof (rule islimptI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3409 | fix T | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3410 | assume "open T" "a \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3411 | from open_right[OF this \<open>a < b\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3412 |   obtain c where c: "a < c" "{a..<c} \<subseteq> T" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3413 | with assms dense[of a "min c b"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3414 |   show "\<exists>y\<in>{a<..<b}. y \<in> T \<and> y \<noteq> a"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3415 | by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3416 | not_le order.strict_implies_order subset_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3417 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3418 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3419 | lemma islimpt_greaterThanLessThan2: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3420 |   fixes a b::"'a::{linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3421 | assumes "a < b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3422 |   shows  "b islimpt {a<..<b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3423 | proof (rule islimptI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3424 | fix T | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3425 | assume "open T" "b \<in> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3426 | from open_left[OF this \<open>a < b\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3427 |   obtain c where c: "c < b" "{c<..b} \<subseteq> T" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3428 | with assms dense[of "max a c" b] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3429 |   show "\<exists>y\<in>{a<..<b}. y \<in> T \<and> y \<noteq> b"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3430 | by (metis greaterThanAtMost_iff greaterThanLessThan_iff max_less_iff_conj | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3431 | not_le order.strict_implies_order subset_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3432 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3433 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3434 | lemma closure_greaterThanLessThan[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3435 |   fixes a b::"'a::{linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3436 |   shows "a < b \<Longrightarrow> closure {a <..< b} = {a .. b}" (is "_ \<Longrightarrow> ?l = ?r")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3437 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3438 | have "?l \<subseteq> closure ?r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3439 | by (rule closure_mono) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3440 |   thus "closure {a<..<b} \<subseteq> {a..b}" by simp
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3441 | qed (auto simp: closure_def order.order_iff_strict islimpt_greaterThanLessThan1 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3442 | islimpt_greaterThanLessThan2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3443 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3444 | lemma closure_greaterThan[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3445 |   fixes a b::"'a::{no_top, linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3446 |   shows "closure {a<..} = {a..}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3447 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3448 | from gt_ex obtain b where "a < b" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3449 |   hence "{a<..} = {a<..<b} \<union> {b..}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3450 |   also have "closure \<dots> = {a..}" using \<open>a < b\<close> unfolding closure_Un
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3451 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3452 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3453 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3454 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3455 | lemma closure_lessThan[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3456 |   fixes b::"'a::{no_bot, linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3457 |   shows "closure {..<b} = {..b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3458 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3459 | from lt_ex obtain a where "a < b" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3460 |   hence "{..<b} = {a<..<b} \<union> {..a}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3461 |   also have "closure \<dots> = {..b}" using \<open>a < b\<close> unfolding closure_Un
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3462 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3463 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3464 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3465 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3466 | lemma closure_atLeastLessThan[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3467 |   fixes a b::"'a::{linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3468 | assumes "a < b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3469 |   shows "closure {a ..< b} = {a .. b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3470 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3471 |   from assms have "{a ..< b} = {a} \<union> {a <..< b}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3472 |   also have "closure \<dots> = {a .. b}" unfolding closure_Un
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3473 | by (auto simp: assms less_imp_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3474 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3475 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3476 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3477 | lemma closure_greaterThanAtMost[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3478 |   fixes a b::"'a::{linorder_topology, dense_order}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3479 | assumes "a < b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3480 |   shows "closure {a <.. b} = {a .. b}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3481 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3482 |   from assms have "{a <.. b} = {b} \<union> {a <..< b}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3483 |   also have "closure \<dots> = {a .. b}" unfolding closure_Un
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3484 | by (auto simp: assms less_imp_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3485 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3486 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3487 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3488 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3489 | subsection \<open>Homeomorphisms\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3490 | |
| 67962 | 3491 | definition%important "homeomorphism s t f g \<longleftrightarrow> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3492 | (\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3493 | (\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3494 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3495 | lemma homeomorphismI [intro?]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3496 | assumes "continuous_on S f" "continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3497 | "f ` S \<subseteq> T" "g ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3498 | shows "homeomorphism S T f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3499 | using assms by (force simp: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3500 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3501 | lemma homeomorphism_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3502 | fixes a :: "'a :: real_normed_vector" | 
| 67399 | 3503 | shows "homeomorphism ((+) a ` S) S ((+) (- a)) ((+) a)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3504 | unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3505 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3506 | lemma homeomorphism_ident: "homeomorphism T T (\<lambda>a. a) (\<lambda>a. a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3507 | by (rule homeomorphismI) (auto simp: continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3508 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3509 | lemma homeomorphism_compose: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3510 | assumes "homeomorphism S T f g" "homeomorphism T U h k" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3511 | shows "homeomorphism S U (h o f) (g o k)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3512 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3513 | unfolding homeomorphism_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3514 | by (intro conjI ballI continuous_on_compose) (auto simp: image_comp [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3515 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3516 | lemma homeomorphism_symD: "homeomorphism S t f g \<Longrightarrow> homeomorphism t S g f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3517 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3518 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3519 | lemma homeomorphism_sym: "homeomorphism S t f g = homeomorphism t S g f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3520 | by (force simp: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3521 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3522 | definition homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3523 | (infixr "homeomorphic" 60) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3524 | where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3525 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3526 | lemma homeomorphic_empty [iff]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3527 |      "S homeomorphic {} \<longleftrightarrow> S = {}" "{} homeomorphic S \<longleftrightarrow> S = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3528 | by (auto simp: homeomorphic_def homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3529 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3530 | lemma homeomorphic_refl: "s homeomorphic s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3531 | unfolding homeomorphic_def homeomorphism_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3532 | using continuous_on_id | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3533 | apply (rule_tac x = "(\<lambda>x. x)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3534 | apply (rule_tac x = "(\<lambda>x. x)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3535 | apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3536 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3537 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3538 | lemma homeomorphic_sym: "s homeomorphic t \<longleftrightarrow> t homeomorphic s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3539 | unfolding homeomorphic_def homeomorphism_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3540 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3541 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3542 | lemma homeomorphic_trans [trans]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3543 | assumes "S homeomorphic T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3544 | and "T homeomorphic U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3545 | shows "S homeomorphic U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3546 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3547 | unfolding homeomorphic_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3548 | by (metis homeomorphism_compose) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3549 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3550 | lemma homeomorphic_minimal: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3551 | "s homeomorphic t \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3552 | (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3553 | (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3554 | continuous_on s f \<and> continuous_on t g)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3555 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3556 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3557 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3558 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3559 | by (fastforce simp: homeomorphic_def homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3560 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3561 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3562 | then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3563 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3564 | unfolding homeomorphic_def homeomorphism_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3565 | by (metis equalityI image_subset_iff subsetI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3566 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3567 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3568 | lemma homeomorphicI [intro?]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3569 | "\<lbrakk>f ` S = T; g ` T = S; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3570 | continuous_on S f; continuous_on T g; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3571 | \<And>x. x \<in> S \<Longrightarrow> g(f(x)) = x; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3572 | \<And>y. y \<in> T \<Longrightarrow> f(g(y)) = y\<rbrakk> \<Longrightarrow> S homeomorphic T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3573 | unfolding homeomorphic_def homeomorphism_def by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3574 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3575 | lemma homeomorphism_of_subsets: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3576 | "\<lbrakk>homeomorphism S T f g; S' \<subseteq> S; T'' \<subseteq> T; f ` S' = T'\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3577 | \<Longrightarrow> homeomorphism S' T' f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3578 | apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3579 | by (metis subsetD imageI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3580 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3581 | lemma homeomorphism_apply1: "\<lbrakk>homeomorphism S T f g; x \<in> S\<rbrakk> \<Longrightarrow> g(f x) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3582 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3583 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3584 | lemma homeomorphism_apply2: "\<lbrakk>homeomorphism S T f g; x \<in> T\<rbrakk> \<Longrightarrow> f(g x) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3585 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3586 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3587 | lemma homeomorphism_image1: "homeomorphism S T f g \<Longrightarrow> f ` S = T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3588 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3589 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3590 | lemma homeomorphism_image2: "homeomorphism S T f g \<Longrightarrow> g ` T = S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3591 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3592 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3593 | lemma homeomorphism_cont1: "homeomorphism S T f g \<Longrightarrow> continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3594 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3595 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3596 | lemma homeomorphism_cont2: "homeomorphism S T f g \<Longrightarrow> continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3597 | by (simp add: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3598 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3599 | lemma continuous_on_no_limpt: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3600 | "(\<And>x. \<not> x islimpt S) \<Longrightarrow> continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3601 | unfolding continuous_on_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3602 | by (metis UNIV_I empty_iff eventually_at_topological islimptE open_UNIV tendsto_def trivial_limit_within) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3603 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3604 | lemma continuous_on_finite: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3605 | fixes S :: "'a::t1_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3606 | shows "finite S \<Longrightarrow> continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3607 | by (metis continuous_on_no_limpt islimpt_finite) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3608 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3609 | lemma homeomorphic_finite: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3610 | fixes S :: "'a::t1_space set" and T :: "'b::t1_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3611 | assumes "finite T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3612 | shows "S homeomorphic T \<longleftrightarrow> finite S \<and> finite T \<and> card S = card T" (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3613 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3614 | assume "S homeomorphic T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3615 | with assms show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3616 | apply (auto simp: homeomorphic_def homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3617 | apply (metis finite_imageI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3618 | by (metis card_image_le finite_imageI le_antisym) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3619 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3620 | assume R: ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3621 | with finite_same_card_bij obtain h where "bij_betw h S T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3622 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3623 | with R show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3624 | apply (auto simp: homeomorphic_def homeomorphism_def continuous_on_finite) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3625 | apply (rule_tac x=h in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3626 | apply (rule_tac x="inv_into S h" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3627 | apply (auto simp: bij_betw_inv_into_left bij_betw_inv_into_right bij_betw_imp_surj_on inv_into_into bij_betwE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3628 | apply (metis bij_betw_def bij_betw_inv_into) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3629 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3630 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3631 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3632 | text \<open>Relatively weak hypotheses if a set is compact.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3633 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3634 | lemma homeomorphism_compact: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3635 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3636 | assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3637 | shows "\<exists>g. homeomorphism s t f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3638 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3639 | define g where "g x = (SOME y. y\<in>s \<and> f y = x)" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3640 | have g: "\<forall>x\<in>s. g (f x) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3641 | using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3642 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3643 | fix y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3644 | assume "y \<in> t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3645 | then obtain x where x:"f x = y" "x\<in>s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3646 | using assms(3) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3647 | then have "g (f x) = x" using g by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3648 | then have "f (g y) = y" unfolding x(1)[symmetric] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3649 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3650 | then have g':"\<forall>x\<in>t. f (g x) = x" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3651 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3652 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3653 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3654 | have "x\<in>s \<Longrightarrow> x \<in> g ` t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3655 | using g[THEN bspec[where x=x]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3656 | unfolding image_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3657 | using assms(3) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3658 | by (auto intro!: bexI[where x="f x"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3659 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3660 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3661 | assume "x\<in>g ` t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3662 | then obtain y where y:"y\<in>t" "g y = x" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3663 | then obtain x' where x':"x'\<in>s" "f x' = y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3664 | using assms(3) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3665 | then have "x \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3666 | unfolding g_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3667 | using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3668 | unfolding y(2)[symmetric] and g_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3669 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3670 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3671 | ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3672 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3673 | then have "g ` t = s" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3674 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3675 | unfolding homeomorphism_def homeomorphic_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3676 | apply (rule_tac x=g in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3677 | using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3678 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3679 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3680 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3681 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3682 | lemma homeomorphic_compact: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3683 | fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3684 | shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s \<Longrightarrow> s homeomorphic t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3685 | unfolding homeomorphic_def by (metis homeomorphism_compact) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3686 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3687 | text\<open>Preservation of topological properties.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3688 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3689 | lemma homeomorphic_compactness: "s homeomorphic t \<Longrightarrow> (compact s \<longleftrightarrow> compact t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3690 | unfolding homeomorphic_def homeomorphism_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3691 | by (metis compact_continuous_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3692 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3693 | text\<open>Results on translation, scaling etc.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3694 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3695 | lemma homeomorphic_scaling: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3696 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3697 | assumes "c \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3698 | shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3699 | unfolding homeomorphic_minimal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3700 | apply (rule_tac x="\<lambda>x. c *\<^sub>R x" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3701 | apply (rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3702 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3703 | apply (auto simp: continuous_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3704 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3705 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3706 | lemma homeomorphic_translation: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3707 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3708 | shows "s homeomorphic ((\<lambda>x. a + x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3709 | unfolding homeomorphic_minimal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3710 | apply (rule_tac x="\<lambda>x. a + x" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3711 | apply (rule_tac x="\<lambda>x. -a + x" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3712 | using continuous_on_add [OF continuous_on_const continuous_on_id, of s a] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3713 | continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3714 | apply auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3715 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3716 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3717 | lemma homeomorphic_affinity: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3718 | fixes s :: "'a::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3719 | assumes "c \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3720 | shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3721 | proof - | 
| 67399 | 3722 | have *: "(+) a ` ( *\<^sub>R) c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3723 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3724 | using homeomorphic_trans | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3725 | using homeomorphic_scaling[OF assms, of s] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3726 | using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3727 | unfolding * | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3728 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3729 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3730 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3731 | lemma homeomorphic_balls: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3732 | fixes a b ::"'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3733 | assumes "0 < d" "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3734 | shows "(ball a d) homeomorphic (ball b e)" (is ?th) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3735 | and "(cball a d) homeomorphic (cball b e)" (is ?cth) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3736 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3737 | show ?th unfolding homeomorphic_minimal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3738 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3739 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3740 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3741 | apply (auto intro!: continuous_intros | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3742 | simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3743 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3744 | show ?cth unfolding homeomorphic_minimal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3745 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3746 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3747 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3748 | apply (auto intro!: continuous_intros | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3749 | simp: dist_commute dist_norm pos_divide_le_eq mult_strict_left_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3750 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3751 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3752 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3753 | lemma homeomorphic_spheres: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3754 | fixes a b ::"'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3755 | assumes "0 < d" "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3756 | shows "(sphere a d) homeomorphic (sphere b e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3757 | unfolding homeomorphic_minimal | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3758 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3759 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3760 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3761 | apply (auto intro!: continuous_intros | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3762 | simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3763 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3764 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3765 | lemma homeomorphic_ball01_UNIV: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3766 | "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3767 | (is "?B homeomorphic ?U") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3768 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3769 | have "x \<in> (\<lambda>z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3770 | apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3771 | apply (auto simp: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3772 | using norm_ge_zero [of x] apply linarith+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3773 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3774 | then show "(\<lambda>z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3775 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3776 | have "x \<in> range (\<lambda>z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3777 | apply (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3778 | using that apply (auto simp: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3779 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3780 | then show "(\<lambda>z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3781 | by (force simp: divide_simps dest: add_less_zeroD) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3782 | show "continuous_on (ball 0 1) (\<lambda>z. z /\<^sub>R (1 - norm z))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3783 | by (rule continuous_intros | force)+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3784 | show "continuous_on UNIV (\<lambda>z. z /\<^sub>R (1 + norm z))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3785 | apply (intro continuous_intros) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3786 | apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3787 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3788 | show "\<And>x. x \<in> ball 0 1 \<Longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3789 | x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3790 | by (auto simp: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3791 | show "\<And>y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3792 | apply (auto simp: divide_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3793 | apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3794 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3795 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3796 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3797 | proposition homeomorphic_ball_UNIV: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3798 | fixes a ::"'a::real_normed_vector" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3799 | assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3800 | using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3801 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3802 | |
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3803 | text \<open>Connectedness is invariant under homeomorphisms.\<close> | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3804 | |
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3805 | lemma homeomorphic_connectedness: | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3806 | assumes "s homeomorphic t" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3807 | shows "connected s \<longleftrightarrow> connected t" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3808 | using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image) | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3809 | |
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67707diff
changeset | 3810 | |
| 67962 | 3811 | subsection%unimportant\<open>Inverse function property for open/closed maps\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3812 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3813 | lemma continuous_on_inverse_open_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3814 | assumes contf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3815 | and imf: "f ` S = T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3816 | and injf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3817 | and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3818 | shows "continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3819 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3820 | from imf injf have gTS: "g ` T = S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3821 | by force | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3822 | from imf injf have fU: "U \<subseteq> S \<Longrightarrow> (f ` U) = T \<inter> g -` U" for U | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3823 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3824 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3825 | by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3826 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3827 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3828 | lemma continuous_on_inverse_closed_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3829 | assumes contf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3830 | and imf: "f ` S = T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3831 | and injf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3832 | and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3833 | shows "continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3834 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3835 | from imf injf have gTS: "g ` T = S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3836 | by force | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3837 | from imf injf have fU: "U \<subseteq> S \<Longrightarrow> (f ` U) = T \<inter> g -` U" for U | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3838 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3839 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3840 | by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3841 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3842 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3843 | lemma homeomorphism_injective_open_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3844 | assumes contf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3845 | and imf: "f ` S = T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3846 | and injf: "inj_on f S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3847 | and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3848 | obtains g where "homeomorphism S T f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3849 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3850 | have "continuous_on T (inv_into S f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3851 | by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3852 | with imf injf contf show "homeomorphism S T f (inv_into S f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3853 | by (auto simp: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3854 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3855 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3856 | lemma homeomorphism_injective_closed_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3857 | assumes contf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3858 | and imf: "f ` S = T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3859 | and injf: "inj_on f S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3860 | and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3861 | obtains g where "homeomorphism S T f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3862 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3863 | have "continuous_on T (inv_into S f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3864 | by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3865 | with imf injf contf show "homeomorphism S T f (inv_into S f)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3866 | by (auto simp: homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3867 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3868 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3869 | lemma homeomorphism_imp_open_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3870 | assumes hom: "homeomorphism S T f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3871 | and oo: "openin (subtopology euclidean S) U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3872 | shows "openin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3873 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3874 | from hom oo have [simp]: "f ` U = T \<inter> g -` U" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3875 | using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3876 | from hom have "continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3877 | unfolding homeomorphism_def by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3878 | moreover have "g ` T = S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3879 | by (metis hom homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3880 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3881 | by (simp add: continuous_on_open oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3882 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3883 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3884 | lemma homeomorphism_imp_closed_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3885 | assumes hom: "homeomorphism S T f g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3886 | and oo: "closedin (subtopology euclidean S) U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3887 | shows "closedin (subtopology euclidean T) (f ` U)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3888 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 3889 | from hom oo have [simp]: "f ` U = T \<inter> g -` U" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3890 | using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3891 | from hom have "continuous_on T g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3892 | unfolding homeomorphism_def by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3893 | moreover have "g ` T = S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3894 | by (metis hom homeomorphism_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3895 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3896 | by (simp add: continuous_on_closed oo) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3897 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3898 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3899 | |
| 67968 | 3900 | subsection \<open>"Isometry" (up to constant bounds) of injective linear map etc\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3901 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3902 | lemma cauchy_isometric: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3903 | assumes e: "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3904 | and s: "subspace s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3905 | and f: "bounded_linear f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3906 | and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3907 | and xs: "\<forall>n. x n \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3908 | and cf: "Cauchy (f \<circ> x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3909 | shows "Cauchy x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3910 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3911 | interpret f: bounded_linear f by fact | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3912 | have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" if "d > 0" for d :: real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3913 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3914 | from that obtain N where N: "\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3915 | using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3916 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3917 | have "norm (x n - x N) < d" if "n \<ge> N" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3918 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3919 | have "e * norm (x n - x N) \<le> norm (f (x n - x N))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3920 | using subspace_diff[OF s, of "x n" "x N"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3921 | using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3922 | using normf[THEN bspec[where x="x n - x N"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3923 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3924 | also have "norm (f (x n - x N)) < e * d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3925 | using \<open>N \<le> n\<close> N unfolding f.diff[symmetric] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3926 | finally show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3927 | using \<open>e>0\<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3928 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3929 | then show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3930 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3931 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3932 | by (simp add: Cauchy_altdef2 dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3933 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3934 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3935 | lemma complete_isometric_image: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3936 | assumes "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3937 | and s: "subspace s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3938 | and f: "bounded_linear f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3939 | and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3940 | and cs: "complete s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3941 | shows "complete (f ` s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3942 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3943 | have "\<exists>l\<in>f ` s. (g \<longlongrightarrow> l) sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3944 | if as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g" for g | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3945 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3946 | from that obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3947 | using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3948 | then have x: "\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3949 | then have "f \<circ> x = g" by (simp add: fun_eq_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3950 | then obtain l where "l\<in>s" and l:"(x \<longlongrightarrow> l) sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3951 | using cs[unfolded complete_def, THEN spec[where x=x]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3952 | using cauchy_isometric[OF \<open>0 < e\<close> s f normf] and cfg and x(1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3953 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3954 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3955 | using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3956 | by (auto simp: \<open>f \<circ> x = g\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3957 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3958 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3959 | unfolding complete_def by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3960 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3961 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 3962 | proposition injective_imp_isometric: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3963 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3964 | assumes s: "closed s" "subspace s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3965 | and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3966 | shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 3967 | proof (cases "s \<subseteq> {0::'a}")
 | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3968 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3969 | have "norm x \<le> norm (f x)" if "x \<in> s" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3970 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3971 | from True that have "x = 0" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3972 | then show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3973 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3974 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3975 | by (auto intro!: exI[where x=1]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3976 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3977 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3978 | interpret f: bounded_linear f by fact | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3979 | from False obtain a where a: "a \<noteq> 0" "a \<in> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3980 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3981 |   from False have "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3982 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3983 |   let ?S = "{f x| x. x \<in> s \<and> norm x = norm a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3984 |   let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3985 |   let ?S'' = "{x::'a. norm x = norm a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3986 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3987 | have "?S'' = frontier (cball 0 (norm a))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3988 | by (simp add: sphere_def dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3989 | then have "compact ?S''" by (metis compact_cball compact_frontier) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3990 | moreover have "?S' = s \<inter> ?S''" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3991 | ultimately have "compact ?S'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3992 | using closed_Int_compact[of s ?S''] using s(1) by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3993 | moreover have *:"f ` ?S' = ?S" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3994 | ultimately have "compact ?S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3995 | using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3996 | then have "closed ?S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3997 | using compact_imp_closed by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3998 |   moreover from a have "?S \<noteq> {}" by auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3999 | ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4000 | using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4001 | then obtain b where "b\<in>s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4002 | and ba: "norm b = norm a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4003 |     and b: "\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4004 | unfolding *[symmetric] unfolding image_iff by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4005 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4006 | let ?e = "norm (f b) / norm b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4007 | have "norm b > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4008 | using ba and a and norm_ge_zero by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4009 | moreover have "norm (f b) > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4010 | using f(2)[THEN bspec[where x=b], OF \<open>b\<in>s\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4011 | using \<open>norm b >0\<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4012 | ultimately have "0 < norm (f b) / norm b" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4013 | moreover | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4014 | have "norm (f b) / norm b * norm x \<le> norm (f x)" if "x\<in>s" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4015 | proof (cases "x = 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4016 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4017 | then show "norm (f b) / norm b * norm x \<le> norm (f x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4018 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4019 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4020 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4021 | with \<open>a \<noteq> 0\<close> have *: "0 < norm a / norm x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4022 | unfolding zero_less_norm_iff[symmetric] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4023 | have "\<forall>x\<in>s. c *\<^sub>R x \<in> s" for c | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4024 | using s[unfolded subspace_def] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4025 |     with \<open>x \<in> s\<close> \<open>x \<noteq> 0\<close> have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4026 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4027 | with \<open>x \<noteq> 0\<close> \<open>a \<noteq> 0\<close> show "norm (f b) / norm b * norm x \<le> norm (f x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4028 | using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4029 | unfolding f.scaleR and ba | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4030 | by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4031 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4032 | ultimately show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4033 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4034 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4035 | proposition closed_injective_image_subspace: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4036 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4037 | assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" "closed s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4038 | shows "closed(f ` s)" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4039 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4040 | obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4041 | using injective_imp_isometric[OF assms(4,1,2,3)] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4042 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4043 | using complete_isometric_image[OF \<open>e>0\<close> assms(1,2) e] and assms(4) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4044 | unfolding complete_eq_closed[symmetric] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4045 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4046 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4047 | |
| 67962 | 4048 | subsection%unimportant \<open>Some properties of a canonical subspace\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4049 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4050 | lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4051 | by (auto simp: subspace_def inner_add_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4052 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4053 | lemma closed_substandard: "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4054 | (is "closed ?A") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4055 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4056 |   let ?D = "{i\<in>Basis. P i}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4057 |   have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4058 | by (simp add: closed_INT closed_Collect_eq continuous_on_inner | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4059 | continuous_on_const continuous_on_id) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4060 |   also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4061 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4062 | finally show "closed ?A" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4063 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4064 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4065 | lemma dim_substandard: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4066 | assumes d: "d \<subseteq> Basis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4067 |   shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4068 | proof (rule dim_unique) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4069 | from d show "d \<subseteq> ?A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4070 | by (auto simp: inner_Basis) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4071 | from d show "independent d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4072 | by (rule independent_mono [OF independent_Basis]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4073 | have "x \<in> span d" if "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4074 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4075 | have "finite d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4076 | by (rule finite_subset [OF d finite_Basis]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4077 | then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4078 | by (simp add: span_sum span_clauses) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4079 | also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4080 | by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4081 | finally show "x \<in> span d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4082 | by (simp only: euclidean_representation) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4083 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4084 | then show "?A \<subseteq> span d" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4085 | qed simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4086 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4087 | text \<open>Hence closure and completeness of all subspaces.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4088 | lemma ex_card: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4089 | assumes "n \<le> card A" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4090 | shows "\<exists>S\<subseteq>A. card S = n" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4091 | proof (cases "finite A") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4092 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4093 |   from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..<card A} A" ..
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4094 |   moreover from f \<open>n \<le> card A\<close> have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4095 | by (auto simp: bij_betw_def intro: subset_inj_on) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4096 |   ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4097 | by (auto simp: bij_betw_def card_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4098 | then show ?thesis by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4099 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4100 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4101 | with \<open>n \<le> card A\<close> show ?thesis by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4102 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4103 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4104 | lemma closed_subspace: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4105 | fixes s :: "'a::euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4106 | assumes "subspace s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4107 | shows "closed s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4108 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4109 | have "dim s \<le> card (Basis :: 'a set)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4110 | using dim_subset_UNIV by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4111 | with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4112 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4113 |   let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4114 |   have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4115 |       inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4116 | using dim_substandard[of d] t d assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4117 | by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4118 | then obtain f where f: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4119 | "linear f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4120 |       "f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4121 |       "inj_on f {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4122 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4123 | interpret f: bounded_linear f | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4124 | using f by (simp add: linear_conv_bounded_linear) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4125 | have "x \<in> ?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4126 | using f.zero d f(3)[THEN inj_onD, of x 0] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4127 | moreover have "closed ?t" by (rule closed_substandard) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4128 | moreover have "subspace ?t" by (rule subspace_substandard) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4129 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4130 | using closed_injective_image_subspace[of ?t f] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4131 | unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4132 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4133 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4134 | lemma complete_subspace: "subspace s \<Longrightarrow> complete s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4135 | for s :: "'a::euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4136 | using complete_eq_closed closed_subspace by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4137 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4138 | lemma closed_span [iff]: "closed (span s)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4139 | for s :: "'a::euclidean_space set" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67968diff
changeset | 4140 | by (simp add: closed_subspace subspace_span) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4141 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4142 | lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4143 | for s :: "'a::euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4144 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4145 | have "?dc \<le> ?d" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67968diff
changeset | 4146 | using closure_minimal[OF span_superset, of s] | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4147 | using closed_subspace[OF subspace_span, of s] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4148 | using dim_subset[of "closure s" "span s"] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4149 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4150 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4151 | using dim_subset[OF closure_subset, of s] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4152 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4153 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4154 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4155 | |
| 67962 | 4156 | subsection%unimportant \<open>Affine transformations of intervals\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4157 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4158 | lemma real_affinity_le: "0 < m \<Longrightarrow> m * x + c \<le> y \<longleftrightarrow> x \<le> inverse m * y + - (c / m)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4159 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4160 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4161 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4162 | lemma real_le_affinity: "0 < m \<Longrightarrow> y \<le> m * x + c \<longleftrightarrow> inverse m * y + - (c / m) \<le> x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4163 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4164 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4165 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4166 | lemma real_affinity_lt: "0 < m \<Longrightarrow> m * x + c < y \<longleftrightarrow> x < inverse m * y + - (c / m)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4167 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4168 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4169 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4170 | lemma real_lt_affinity: "0 < m \<Longrightarrow> y < m * x + c \<longleftrightarrow> inverse m * y + - (c / m) < x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4171 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4172 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4173 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4174 | lemma real_affinity_eq: "m \<noteq> 0 \<Longrightarrow> m * x + c = y \<longleftrightarrow> x = inverse m * y + - (c / m)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4175 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4176 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4177 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4178 | lemma real_eq_affinity: "m \<noteq> 0 \<Longrightarrow> y = m * x + c \<longleftrightarrow> inverse m * y + - (c / m) = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4179 | for m :: "'a::linordered_field" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4180 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4181 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4182 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4183 | subsection \<open>Banach fixed point theorem (not really topological ...)\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4184 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4185 | theorem banach_fix: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4186 |   assumes s: "complete s" "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4187 | and c: "0 \<le> c" "c < 1" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4188 | and f: "f ` s \<subseteq> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4189 | and lipschitz: "\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4190 | shows "\<exists>!x\<in>s. f x = x" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4191 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4192 | from c have "1 - c > 0" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4193 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4194 | from s(2) obtain z0 where z0: "z0 \<in> s" by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4195 | define z where "z n = (f ^^ n) z0" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4196 | with f z0 have z_in_s: "z n \<in> s" for n :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4197 | by (induct n) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4198 | define d where "d = dist (z 0) (z 1)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4199 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4200 | have fzn: "f (z n) = z (Suc n)" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4201 | by (simp add: z_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4202 | have cf_z: "dist (z n) (z (Suc n)) \<le> (c ^ n) * d" for n :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4203 | proof (induct n) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4204 | case 0 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4205 | then show ?case | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4206 | by (simp add: d_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4207 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4208 | case (Suc m) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4209 | with \<open>0 \<le> c\<close> have "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4210 | using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4211 | then show ?case | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4212 | using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4213 | by (simp add: fzn mult_le_cancel_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4214 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4215 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4216 | have cf_z2: "(1 - c) * dist (z m) (z (m + n)) \<le> (c ^ m) * d * (1 - c ^ n)" for n m :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4217 | proof (induct n) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4218 | case 0 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4219 | show ?case by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4220 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4221 | case (Suc k) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4222 | from c have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4223 | (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4224 | by (simp add: dist_triangle) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4225 | also from c cf_z[of "m + k"] have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4226 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4227 | also from Suc have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4228 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4229 | also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4230 | by (simp add: power_add field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4231 | also from c have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4232 | by (simp add: field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4233 | finally show ?case by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4234 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4235 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4236 | have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e" if "e > 0" for e | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4237 | proof (cases "d = 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4238 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4239 | from \<open>1 - c > 0\<close> have "(1 - c) * x \<le> 0 \<longleftrightarrow> x \<le> 0" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4240 | by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4241 | with c cf_z2[of 0] True have "z n = z0" for n | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4242 | by (simp add: z_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4243 | with \<open>e > 0\<close> show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4244 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4245 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4246 | with zero_le_dist[of "z 0" "z 1"] have "d > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4247 | by (metis d_def less_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4248 | with \<open>1 - c > 0\<close> \<open>e > 0\<close> have "0 < e * (1 - c) / d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4249 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4250 | with c obtain N where N: "c ^ N < e * (1 - c) / d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4251 | using real_arch_pow_inv[of "e * (1 - c) / d" c] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4252 | have *: "dist (z m) (z n) < e" if "m > n" and as: "m \<ge> N" "n \<ge> N" for m n :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4253 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4254 | from c \<open>n \<ge> N\<close> have *: "c ^ n \<le> c ^ N" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4255 | using power_decreasing[OF \<open>n\<ge>N\<close>, of c] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4256 | from c \<open>m > n\<close> have "1 - c ^ (m - n) > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4257 | using power_strict_mono[of c 1 "m - n"] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4258 | with \<open>d > 0\<close> \<open>0 < 1 - c\<close> have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4259 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4260 | from cf_z2[of n "m - n"] \<open>m > n\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4261 | have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4262 | by (simp add: pos_le_divide_eq[OF \<open>1 - c > 0\<close>] mult.commute dist_commute) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4263 | also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4264 | using mult_right_mono[OF * order_less_imp_le[OF **]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4265 | by (simp add: mult.assoc) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4266 | also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4267 | using mult_strict_right_mono[OF N **] by (auto simp: mult.assoc) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4268 | also from c \<open>d > 0\<close> \<open>1 - c > 0\<close> have "\<dots> = e * (1 - c ^ (m - n))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4269 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4270 | also from c \<open>1 - c ^ (m - n) > 0\<close> \<open>e > 0\<close> have "\<dots> \<le> e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4271 | using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4272 | finally show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4273 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4274 | have "dist (z n) (z m) < e" if "N \<le> m" "N \<le> n" for m n :: nat | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4275 | proof (cases "n = m") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4276 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4277 | with \<open>e > 0\<close> show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4278 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4279 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4280 | with *[of n m] *[of m n] and that show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4281 | by (auto simp: dist_commute nat_neq_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4282 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4283 | then show ?thesis by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4284 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4285 | then have "Cauchy z" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4286 | by (simp add: cauchy_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4287 | then obtain x where "x\<in>s" and x:"(z \<longlongrightarrow> x) sequentially" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4288 | using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4289 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4290 | define e where "e = dist (f x) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4291 | have "e = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4292 | proof (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4293 | assume "e \<noteq> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4294 | then have "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4295 | unfolding e_def using zero_le_dist[of "f x" x] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4296 | by (metis dist_eq_0_iff dist_nz e_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4297 | then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4298 | using x[unfolded lim_sequentially, THEN spec[where x="e/2"]] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4299 | then have N':"dist (z N) x < e / 2" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4300 | have *: "c * dist (z N) x \<le> dist (z N) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4301 | unfolding mult_le_cancel_right2 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4302 | using zero_le_dist[of "z N" x] and c | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4303 | by (metis dist_eq_0_iff dist_nz order_less_asym less_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4304 | have "dist (f (z N)) (f x) \<le> c * dist (z N) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4305 | using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4306 | using z_in_s[of N] \<open>x\<in>s\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4307 | using c | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4308 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4309 | also have "\<dots> < e / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4310 | using N' and c using * by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4311 | finally show False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4312 | unfolding fzn | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4313 | using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4314 | unfolding e_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4315 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4316 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4317 | then have "f x = x" by (auto simp: e_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4318 | moreover have "y = x" if "f y = y" "y \<in> s" for y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4319 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4320 | from \<open>x \<in> s\<close> \<open>f x = x\<close> that have "dist x y \<le> c * dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4321 | using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4322 | with c and zero_le_dist[of x y] have "dist x y = 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4323 | by (simp add: mult_le_cancel_right1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4324 | then show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4325 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4326 | ultimately show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4327 | using \<open>x\<in>s\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4328 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4329 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4330 | lemma banach_fix_type: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4331 | fixes f::"'a::complete_space\<Rightarrow>'a" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4332 | assumes c:"0 \<le> c" "c < 1" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4333 | and lipschitz:"\<forall>x. \<forall>y. dist (f x) (f y) \<le> c * dist x y" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4334 | shows "\<exists>!x. (f x = x)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4335 | using assms banach_fix[OF complete_UNIV UNIV_not_empty assms(1,2) subset_UNIV, of f] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4336 | by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 4337 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4338 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4339 | subsection \<open>Edelstein fixed point theorem\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4340 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4341 | theorem edelstein_fix: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4342 | fixes s :: "'a::metric_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4343 |   assumes s: "compact s" "s \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4344 | and gs: "(g ` s) \<subseteq> s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4345 | and dist: "\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4346 | shows "\<exists>!x\<in>s. g x = x" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4347 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4348 | let ?D = "(\<lambda>x. (x, x)) ` s" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4349 |   have D: "compact ?D" "?D \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4350 | by (rule compact_continuous_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4351 | (auto intro!: s continuous_Pair continuous_ident simp: continuous_on_eq_continuous_within) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4352 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4353 | have "\<And>x y e. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> 0 < e \<Longrightarrow> dist y x < e \<Longrightarrow> dist (g y) (g x) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4354 | using dist by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4355 | then have "continuous_on s g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4356 | by (auto simp: continuous_on_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4357 | then have cont: "continuous_on ?D (\<lambda>x. dist ((g \<circ> fst) x) (snd x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4358 | unfolding continuous_on_eq_continuous_within | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4359 | by (intro continuous_dist ballI continuous_within_compose) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4360 | (auto intro!: continuous_fst continuous_snd continuous_ident simp: image_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4361 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4362 | obtain a where "a \<in> s" and le: "\<And>x. x \<in> s \<Longrightarrow> dist (g a) a \<le> dist (g x) x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4363 | using continuous_attains_inf[OF D cont] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4364 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4365 | have "g a = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4366 | proof (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4367 | assume "g a \<noteq> a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4368 | with \<open>a \<in> s\<close> gs have "dist (g (g a)) (g a) < dist (g a) a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4369 | by (intro dist[rule_format]) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4370 | moreover have "dist (g a) a \<le> dist (g (g a)) (g a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4371 | using \<open>a \<in> s\<close> gs by (intro le) auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4372 | ultimately show False by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4373 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4374 | moreover have "\<And>x. x \<in> s \<Longrightarrow> g x = x \<Longrightarrow> x = a" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4375 | using dist[THEN bspec[where x=a]] \<open>g a = a\<close> and \<open>a\<in>s\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4376 | ultimately show "\<exists>!x\<in>s. g x = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4377 | using \<open>a \<in> s\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4378 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4379 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4380 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4381 | lemma cball_subset_cball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4382 | fixes a :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4383 | shows "cball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r < 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4384 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4385 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4386 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4387 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4388 | proof (cases "r < 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4389 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4390 | then show ?rhs by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4391 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4392 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4393 | then have [simp]: "r \<ge> 0" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4394 | have "norm (a - a') + r \<le> r'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4395 | proof (cases "a = a'") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4396 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4397 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4398 | using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4399 | by (force simp: SOME_Basis dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4400 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4401 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4402 | have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4403 | by (simp add: algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4404 | also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4405 | by (simp add: algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4406 | also from \<open>a \<noteq> a'\<close> have "... = \<bar>- norm (a - a') - r\<bar>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4407 | by (simp add: abs_mult_pos field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4408 | finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \<bar>norm (a - a') + r\<bar>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4409 | by linarith | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4410 | from \<open>a \<noteq> a'\<close> show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4411 | using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4412 | by (simp add: dist_norm scaleR_add_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4413 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4414 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4415 | by (simp add: dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4416 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4417 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4418 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4419 | then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4420 | by (auto simp: ball_def dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4421 | (metis add.commute add_le_cancel_right dist_norm dist_triangle3 order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4422 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4423 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4424 | lemma cball_subset_ball_iff: "cball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r < r' \<or> r < 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4425 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4426 | for a :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4427 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4428 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4429 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4430 | proof (cases "r < 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4431 | case True then | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4432 | show ?rhs by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4433 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4434 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4435 | then have [simp]: "r \<ge> 0" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4436 | have "norm (a - a') + r < r'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4437 | proof (cases "a = a'") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4438 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4439 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4440 | using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4441 | by (force simp: SOME_Basis dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4442 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4443 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4444 | have False if "norm (a - a') + r \<ge> r'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4445 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4446 | from that have "\<bar>r' - norm (a - a')\<bar> \<le> r" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4447 | by (simp split: abs_split) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4448 | (metis \<open>0 \<le> r\<close> \<open>?lhs\<close> centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4449 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4450 | using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] \<open>a \<noteq> a'\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4451 | by (simp add: dist_norm field_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4452 | (simp add: diff_divide_distrib scaleR_left_diff_distrib) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4453 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4454 | then show ?thesis by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4455 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4456 | then show ?rhs by (simp add: dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4457 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4458 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4459 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4460 | then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4461 | by (auto simp: ball_def dist_norm) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4462 | (metis add.commute add_le_cancel_right dist_norm dist_triangle3 le_less_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4463 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4464 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4465 | lemma ball_subset_cball_iff: "ball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4466 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4467 | for a :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4468 | proof (cases "r \<le> 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4469 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4470 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4471 | using dist_not_less_zero less_le_trans by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4472 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4473 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4474 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4475 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4476 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4477 | then have "(cball a r \<subseteq> cball a' r')" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4478 | by (metis False closed_cball closure_ball closure_closed closure_mono not_less) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4479 | with False show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4480 | by (fastforce iff: cball_subset_cball_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4481 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4482 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4483 | with False show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4484 | using ball_subset_cball cball_subset_cball_iff by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4485 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4486 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4487 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4488 | lemma ball_subset_ball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4489 | fixes a :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4490 | shows "ball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4491 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4492 | proof (cases "r \<le> 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4493 | case True then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4494 | using dist_not_less_zero less_le_trans by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4495 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4496 | case False show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4497 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4498 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4499 | then have "0 < r'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4500 | by (metis (no_types) False \<open>?lhs\<close> centre_in_ball dist_norm le_less_trans mem_ball norm_ge_zero not_less set_mp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4501 | then have "(cball a r \<subseteq> cball a' r')" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4502 | by (metis False\<open>?lhs\<close> closure_ball closure_mono not_less) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4503 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4504 | using False cball_subset_cball_iff by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4505 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4506 | assume ?rhs then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4507 | apply (auto simp: ball_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4508 | apply (metis add.commute add_le_cancel_right dist_commute dist_triangle_lt not_le order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4509 | using dist_not_less_zero order.strict_trans2 apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4510 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4511 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4512 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4513 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4514 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4515 | lemma ball_eq_ball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4516 | fixes x :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4517 | shows "ball x d = ball y e \<longleftrightarrow> d \<le> 0 \<and> e \<le> 0 \<or> x=y \<and> d=e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4518 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4519 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4520 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4521 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4522 | proof (cases "d \<le> 0 \<or> e \<le> 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4523 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4524 | with \<open>?lhs\<close> show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4525 | by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4526 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4527 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4528 | with \<open>?lhs\<close> show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4529 | apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4530 | apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4531 | apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4532 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4533 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4534 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4535 | assume ?rhs then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4536 | by (auto simp: set_eq_subset ball_subset_ball_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4537 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4538 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4539 | lemma cball_eq_cball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4540 | fixes x :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4541 | shows "cball x d = cball y e \<longleftrightarrow> d < 0 \<and> e < 0 \<or> x=y \<and> d=e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4542 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4543 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4544 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4545 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4546 | proof (cases "d < 0 \<or> e < 0") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4547 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4548 | with \<open>?lhs\<close> show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4549 | by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4550 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4551 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4552 | with \<open>?lhs\<close> show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4553 | apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4554 | apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4555 | apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4556 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4557 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4558 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4559 | assume ?rhs then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4560 | by (auto simp: set_eq_subset cball_subset_cball_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4561 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4562 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4563 | lemma ball_eq_cball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4564 | fixes x :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4565 | shows "ball x d = cball y e \<longleftrightarrow> d \<le> 0 \<and> e < 0" (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4566 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4567 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4568 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4569 | apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4570 | apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4571 | apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4572 | using \<open>?lhs\<close> ball_eq_empty cball_eq_empty apply blast+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4573 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4574 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4575 | assume ?rhs then show ?lhs by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4576 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4577 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4578 | lemma cball_eq_ball_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4579 | fixes x :: "'a :: euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4580 | shows "cball x d = ball y e \<longleftrightarrow> d < 0 \<and> e \<le> 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4581 | using ball_eq_cball_iff by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4582 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4583 | lemma finite_ball_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4584 | fixes S :: "'a :: euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4585 | assumes "open S" "finite X" "p \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4586 | shows "\<exists>e>0. \<forall>w\<in>ball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4587 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4588 | obtain e1 where "0 < e1" and e1_b:"ball p e1 \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4589 | using open_contains_ball_eq[OF \<open>open S\<close>] assms by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4590 | obtain e2 where "0 < e2" and "\<forall>x\<in>X. x \<noteq> p \<longrightarrow> e2 \<le> dist p x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4591 | using finite_set_avoid[OF \<open>finite X\<close>,of p] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4592 | hence "\<forall>w\<in>ball p (min e1 e2). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" using e1_b by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4593 | thus "\<exists>e>0. \<forall>w\<in>ball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> \<open>e1>0\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4594 | apply (rule_tac x="min e1 e2" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4595 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4596 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4597 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4598 | lemma finite_cball_avoid: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4599 | fixes S :: "'a :: euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4600 | assumes "open S" "finite X" "p \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4601 | shows "\<exists>e>0. \<forall>w\<in>cball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4602 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4603 | obtain e1 where "e1>0" and e1: "\<forall>w\<in>ball p e1. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4604 | using finite_ball_avoid[OF assms] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4605 | define e2 where "e2 \<equiv> e1/2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4606 | have "e2>0" and "e2 < e1" unfolding e2_def using \<open>e1>0\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4607 | then have "cball p e2 \<subseteq> ball p e1" by (subst cball_subset_ball_iff,auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4608 | then show "\<exists>e>0. \<forall>w\<in>cball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> e1 by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4609 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4610 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4611 | subsection\<open>Various separability-type properties\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4612 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4613 | lemma univ_second_countable: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4614 | obtains \<B> :: "'a::euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4615 | where "countable \<B>" "\<And>C. C \<in> \<B> \<Longrightarrow> open C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4616 | "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4617 | by (metis ex_countable_basis topological_basis_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4618 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4619 | lemma subset_second_countable: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4620 | obtains \<B> :: "'a:: euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4621 | where "countable \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4622 |           "{} \<notin> \<B>"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4623 | "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4624 | "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4625 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4626 | obtain \<B> :: "'a set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4627 | where "countable \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4628 | and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4629 | and \<B>: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4630 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4631 | obtain \<C> :: "'a set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4632 | where "countable \<C>" and ope: "\<And>C. C \<in> \<C> \<Longrightarrow> open C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4633 | and \<C>: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<C> \<and> S = \<Union>U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4634 | by (metis univ_second_countable that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4635 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4636 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4637 | show "countable ((\<lambda>C. S \<inter> C) ` \<C>)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4638 | by (simp add: \<open>countable \<C>\<close>) | 
| 67399 | 4639 | show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (subtopology euclidean S) C" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4640 | using ope by auto | 
| 67399 | 4641 | show "\<And>T. openin (subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4642 | by (metis \<C> image_mono inf_Sup openin_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4643 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4644 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4645 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4646 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4647 |     show "countable (\<B> - {{}})"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4648 | using \<open>countable \<B>\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4649 |     show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) C"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4650 | by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (subtopology euclidean S) C\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4651 |     show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (subtopology euclidean S) T" for T
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4652 | using \<B> [OF that] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4653 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4654 |       apply (rule_tac x="\<U> - {{}}" in exI, auto)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4655 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4656 | qed auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4657 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4658 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4659 | lemma univ_second_countable_sequence: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4660 | obtains B :: "nat \<Rightarrow> 'a::euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4661 |     where "inj B" "\<And>n. open(B n)" "\<And>S. open S \<Longrightarrow> \<exists>k. S = \<Union>{B n |n. n \<in> k}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4662 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4663 | obtain \<B> :: "'a set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4664 | where "countable \<B>" | 
| 67237 | 4665 | and opn: "\<And>C. C \<in> \<B> \<Longrightarrow> open C" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4666 | and Un: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4667 | using univ_second_countable by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4668 | have *: "infinite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4669 | apply (rule Infinite_Set.range_inj_infinite) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4670 | apply (simp add: inj_on_def ball_eq_ball_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4671 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4672 | have "infinite \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4673 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4674 | assume "finite \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4675 | then have "finite (Union ` (Pow \<B>))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4676 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4677 | then have "finite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4678 | apply (rule rev_finite_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4679 | by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4680 | with * show False by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4681 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4682 | obtain f :: "nat \<Rightarrow> 'a set" where "\<B> = range f" "inj f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4683 | by (blast intro: countable_as_injective_image [OF \<open>countable \<B>\<close> \<open>infinite \<B>\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4684 |   have *: "\<exists>k. S = \<Union>{f n |n. n \<in> k}" if "open S" for S
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4685 | using Un [OF that] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4686 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4687 | apply (rule_tac x="f-`U" in exI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4688 | using \<open>inj f\<close> \<open>\<B> = range f\<close> apply force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4689 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4690 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4691 | apply (rule that [OF \<open>inj f\<close> _ *]) | 
| 67237 | 4692 | apply (auto simp: \<open>\<B> = range f\<close> opn) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4693 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4694 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4695 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4696 | proposition separable: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4697 | fixes S :: "'a:: euclidean_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4698 | obtains T where "countable T" "T \<subseteq> S" "S \<subseteq> closure T" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4699 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4700 | obtain \<B> :: "'a:: euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4701 | where "countable \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4702 |       and "{} \<notin> \<B>"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4703 | and ope: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4704 | and if_ope: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4705 | by (meson subset_second_countable) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4706 | then obtain f where f: "\<And>C. C \<in> \<B> \<Longrightarrow> f C \<in> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4707 | by (metis equals0I) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4708 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4709 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4710 | show "countable (f ` \<B>)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4711 | by (simp add: \<open>countable \<B>\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4712 | show "f ` \<B> \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4713 | using ope f openin_imp_subset by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4714 | show "S \<subseteq> closure (f ` \<B>)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4715 | proof (clarsimp simp: closure_approachable) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4716 | fix x and e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4717 | assume "x \<in> S" "0 < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4718 | have "openin (subtopology euclidean S) (S \<inter> ball x e)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4719 | by (simp add: openin_Int_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4720 | with if_ope obtain \<U> where \<U>: "\<U> \<subseteq> \<B>" "S \<inter> ball x e = \<Union>\<U>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4721 | by meson | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4722 | show "\<exists>C \<in> \<B>. dist (f C) x < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4723 |       proof (cases "\<U> = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4724 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4725 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4726 | using \<open>0 < e\<close> \<U> \<open>x \<in> S\<close> by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4727 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4728 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4729 | then obtain C where "C \<in> \<U>" by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4730 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4731 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4732 | show "dist (f C) x < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4733 | by (metis Int_iff Union_iff \<U> \<open>C \<in> \<U>\<close> dist_commute f mem_ball subsetCE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4734 | show "C \<in> \<B>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4735 | using \<open>\<U> \<subseteq> \<B>\<close> \<open>C \<in> \<U>\<close> by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4736 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4737 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4738 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4739 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4740 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4741 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4742 | proposition Lindelof: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4743 | fixes \<F> :: "'a::euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4744 | assumes \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> open S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4745 | obtains \<F>' where "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 4746 | proof - | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4747 | obtain \<B> :: "'a set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4748 | where "countable \<B>" "\<And>C. C \<in> \<B> \<Longrightarrow> open C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4749 | and \<B>: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4750 | using univ_second_countable by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4751 |   define \<D> where "\<D> \<equiv> {S. S \<in> \<B> \<and> (\<exists>U. U \<in> \<F> \<and> S \<subseteq> U)}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4752 | have "countable \<D>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4753 | apply (rule countable_subset [OF _ \<open>countable \<B>\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4754 | apply (force simp: \<D>_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4755 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4756 | have "\<And>S. \<exists>U. S \<in> \<D> \<longrightarrow> U \<in> \<F> \<and> S \<subseteq> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4757 | by (simp add: \<D>_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4758 | then obtain G where G: "\<And>S. S \<in> \<D> \<longrightarrow> G S \<in> \<F> \<and> S \<subseteq> G S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4759 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4760 | have "\<Union>\<F> \<subseteq> \<Union>\<D>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4761 | unfolding \<D>_def by (blast dest: \<F> \<B>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4762 | moreover have "\<Union>\<D> \<subseteq> \<Union>\<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4763 | using \<D>_def by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4764 | ultimately have eq1: "\<Union>\<F> = \<Union>\<D>" .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4765 | have eq2: "\<Union>\<D> = UNION \<D> G" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4766 | using G eq1 by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4767 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4768 | apply (rule_tac \<F>' = "G ` \<D>" in that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4769 | using G \<open>countable \<D>\<close> apply (auto simp: eq1 eq2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4770 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4771 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4772 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4773 | lemma Lindelof_openin: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4774 | fixes \<F> :: "'a::euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4775 | assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean U) S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4776 | obtains \<F>' where "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4777 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4778 | have "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>T. open T \<and> S = U \<inter> T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4779 | using assms by (simp add: openin_open) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4780 | then obtain tf where tf: "\<And>S. S \<in> \<F> \<Longrightarrow> open (tf S) \<and> (S = U \<inter> tf S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4781 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4782 | have [simp]: "\<And>\<F>'. \<F>' \<subseteq> \<F> \<Longrightarrow> \<Union>\<F>' = U \<inter> \<Union>(tf ` \<F>')" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4783 | using tf by fastforce | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4784 | obtain \<G> where "countable \<G> \<and> \<G> \<subseteq> tf ` \<F>" "\<Union>\<G> = UNION \<F> tf" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4785 | using tf by (force intro: Lindelof [of "tf ` \<F>"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4786 | then obtain \<F>' where \<F>': "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4787 | by (clarsimp simp add: countable_subset_image) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4788 | then show ?thesis .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4789 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4790 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4791 | lemma countable_disjoint_open_subsets: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4792 | fixes \<F> :: "'a::euclidean_space set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4793 | assumes "\<And>S. S \<in> \<F> \<Longrightarrow> open S" and pw: "pairwise disjnt \<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4794 | shows "countable \<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4795 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4796 | obtain \<F>' where "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4797 | by (meson assms Lindelof) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4798 |   with pw have "\<F> \<subseteq> insert {} \<F>'"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4799 | by (fastforce simp add: pairwise_def disjnt_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4800 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4801 | by (simp add: \<open>countable \<F>'\<close> countable_subset) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4802 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4803 | |
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4804 | lemma countable_disjoint_nonempty_interior_subsets: | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4805 | fixes \<F> :: "'a::euclidean_space set set" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4806 |   assumes pw: "pairwise disjnt \<F>" and int: "\<And>S. \<lbrakk>S \<in> \<F>; interior S = {}\<rbrakk> \<Longrightarrow> S = {}"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4807 | shows "countable \<F>" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4808 | proof (rule countable_image_inj_on) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4809 | have "disjoint (interior ` \<F>)" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4810 | using pw by (simp add: disjoint_image_subset interior_subset) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4811 | then show "countable (interior ` \<F>)" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4812 | by (auto intro: countable_disjoint_open_subsets) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4813 | show "inj_on interior \<F>" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4814 | using pw apply (clarsimp simp: inj_on_def pairwise_def) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4815 | apply (metis disjnt_def disjnt_subset1 inf.orderE int interior_subset) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4816 | done | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4817 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 4818 | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4819 | lemma closedin_compact: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4820 | "\<lbrakk>compact S; closedin (subtopology euclidean S) T\<rbrakk> \<Longrightarrow> compact T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4821 | by (metis closedin_closed compact_Int_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4822 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4823 | lemma closedin_compact_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4824 | fixes S :: "'a::t2_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4825 | shows | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4826 | "compact S | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4827 | \<Longrightarrow> (closedin (subtopology euclidean S) T \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4828 | compact T \<and> T \<subseteq> S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4829 | by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4830 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4831 | lemma continuous_imp_closed_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4832 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4833 | assumes "closedin (subtopology euclidean S) U" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4834 | "continuous_on S f" "f ` S = T" "compact S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4835 | shows "closedin (subtopology euclidean T) (f ` U)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4836 | by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4837 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4838 | lemma continuous_imp_quotient_map: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4839 | fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4840 | assumes "continuous_on S f" "f ` S = T" "compact S" "U \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4841 | shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4842 | openin (subtopology euclidean T) U" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4843 | by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4844 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4845 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4846 | lemma open_map_restrict: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4847 | assumes opeU: "openin (subtopology euclidean (S \<inter> f -` T')) U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4848 | and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4849 | and "T' \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4850 | shows "openin (subtopology euclidean T') (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4851 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4852 | obtain V where "open V" "U = S \<inter> f -` T' \<inter> V" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4853 | using opeU by (auto simp: openin_open) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4854 | with oo [of "S \<inter> V"] \<open>T' \<subseteq> T\<close> show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4855 | by (fastforce simp add: openin_open) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4856 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4857 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4858 | lemma closed_map_restrict: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4859 | assumes cloU: "closedin (subtopology euclidean (S \<inter> f -` T')) U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4860 | and cc: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4861 | and "T' \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4862 | shows "closedin (subtopology euclidean T') (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4863 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4864 | obtain V where "closed V" "U = S \<inter> f -` T' \<inter> V" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4865 | using cloU by (auto simp: closedin_closed) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4866 | with cc [of "S \<inter> V"] \<open>T' \<subseteq> T\<close> show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4867 | by (fastforce simp add: closedin_closed) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4868 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4869 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4870 | lemma connected_monotone_quotient_preimage: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4871 | assumes "connected T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4872 | and contf: "continuous_on S f" and fim: "f ` S = T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4873 | and opT: "\<And>U. U \<subseteq> T | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4874 | \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4875 | openin (subtopology euclidean T) U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4876 |       and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4877 | shows "connected S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4878 | proof (rule connectedI) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4879 | fix U V | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4880 |   assume "open U" and "open V" and "U \<inter> S \<noteq> {}" and "V \<inter> S \<noteq> {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4881 |     and "U \<inter> V \<inter> S = {}" and "S \<subseteq> U \<union> V"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4882 | moreover | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4883 |   have disjoint: "f ` (S \<inter> U) \<inter> f ` (S \<inter> V) = {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4884 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4885 | have False if "y \<in> f ` (S \<inter> U) \<inter> f ` (S \<inter> V)" for y | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4886 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4887 | have "y \<in> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4888 | using fim that by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4889 | show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4890 | using connectedD [OF connT [OF \<open>y \<in> T\<close>] \<open>open U\<close> \<open>open V\<close>] | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4891 |               \<open>S \<subseteq> U \<union> V\<close> \<open>U \<inter> V \<inter> S = {}\<close> that by fastforce
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4892 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4893 | then show ?thesis by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4894 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4895 | ultimately have UU: "(S \<inter> f -` f ` (S \<inter> U)) = S \<inter> U" and VV: "(S \<inter> f -` f ` (S \<inter> V)) = S \<inter> V" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4896 | by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4897 | have opeU: "openin (subtopology euclidean T) (f ` (S \<inter> U))" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4898 | by (metis UU \<open>open U\<close> fim image_Int_subset le_inf_iff opT openin_open_Int) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4899 | have opeV: "openin (subtopology euclidean T) (f ` (S \<inter> V))" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4900 | by (metis opT fim VV \<open>open V\<close> openin_open_Int image_Int_subset inf.bounded_iff) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4901 | have "T \<subseteq> f ` (S \<inter> U) \<union> f ` (S \<inter> V)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4902 | using \<open>S \<subseteq> U \<union> V\<close> fim by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4903 | then show False | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4904 |     using \<open>connected T\<close> disjoint opeU opeV \<open>U \<inter> S \<noteq> {}\<close> \<open>V \<inter> S \<noteq> {}\<close>
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4905 | by (auto simp: connected_openin) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4906 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4907 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4908 | lemma connected_open_monotone_preimage: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4909 | assumes contf: "continuous_on S f" and fim: "f ` S = T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4910 | and ST: "\<And>C. openin (subtopology euclidean S) C \<Longrightarrow> openin (subtopology euclidean T) (f ` C)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4911 |     and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4912 | and "connected C" "C \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4913 | shows "connected (S \<inter> f -` C)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4914 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4915 | have contf': "continuous_on (S \<inter> f -` C) f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4916 | by (meson contf continuous_on_subset inf_le1) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4917 | have eqC: "f ` (S \<inter> f -` C) = C" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4918 | using \<open>C \<subseteq> T\<close> fim by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4919 | show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4920 | proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4921 |     show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4922 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4923 |       have "S \<inter> f -` C \<inter> f -` {y} = S \<inter> f -` {y}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4924 | using that by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4925 |       moreover have "connected (S \<inter> f -` {y})"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4926 | using \<open>C \<subseteq> T\<close> connT that by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4927 | ultimately show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4928 | by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4929 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4930 | have "\<And>U. openin (subtopology euclidean (S \<inter> f -` C)) U | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4931 | \<Longrightarrow> openin (subtopology euclidean C) (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4932 | using open_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4933 | then show "\<And>D. D \<subseteq> C | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4934 | \<Longrightarrow> openin (subtopology euclidean (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) = | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4935 | openin (subtopology euclidean C) D" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4936 | using open_map_imp_quotient_map [of "(S \<inter> f -` C)" f] contf' by (simp add: eqC) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4937 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4938 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4939 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4940 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4941 | lemma connected_closed_monotone_preimage: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4942 | assumes contf: "continuous_on S f" and fim: "f ` S = T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4943 | and ST: "\<And>C. closedin (subtopology euclidean S) C \<Longrightarrow> closedin (subtopology euclidean T) (f ` C)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4944 |     and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4945 | and "connected C" "C \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4946 | shows "connected (S \<inter> f -` C)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4947 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4948 | have contf': "continuous_on (S \<inter> f -` C) f" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4949 | by (meson contf continuous_on_subset inf_le1) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4950 | have eqC: "f ` (S \<inter> f -` C) = C" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4951 | using \<open>C \<subseteq> T\<close> fim by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4952 | show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4953 | proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4954 |     show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4955 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4956 |       have "S \<inter> f -` C \<inter> f -` {y} = S \<inter> f -` {y}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4957 | using that by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4958 |       moreover have "connected (S \<inter> f -` {y})"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4959 | using \<open>C \<subseteq> T\<close> connT that by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4960 | ultimately show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4961 | by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4962 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4963 | have "\<And>U. closedin (subtopology euclidean (S \<inter> f -` C)) U | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4964 | \<Longrightarrow> closedin (subtopology euclidean C) (f ` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4965 | using closed_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4966 | then show "\<And>D. D \<subseteq> C | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4967 | \<Longrightarrow> openin (subtopology euclidean (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) = | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4968 | openin (subtopology euclidean C) D" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4969 | using closed_map_imp_quotient_map [of "(S \<inter> f -` C)" f] contf' by (simp add: eqC) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4970 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4971 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4972 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4973 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4974 | |
| 67968 | 4975 | subsection\<open>A couple of lemmas about components (see Newman IV, 3.3 and 3.4)\<close> | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4976 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4977 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4978 | lemma connected_Un_clopen_in_complement: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4979 | fixes S U :: "'a::metric_space set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4980 | assumes "connected S" "connected U" "S \<subseteq> U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4981 | and opeT: "openin (subtopology euclidean (U - S)) T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4982 | and cloT: "closedin (subtopology euclidean (U - S)) T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4983 | shows "connected (S \<union> T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4984 | proof - | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4985 | have *: "\<lbrakk>\<And>x y. P x y \<longleftrightarrow> P y x; \<And>x y. P x y \<Longrightarrow> S \<subseteq> x \<or> S \<subseteq> y; | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4986 | \<And>x y. \<lbrakk>P x y; S \<subseteq> x\<rbrakk> \<Longrightarrow> False\<rbrakk> \<Longrightarrow> ~(\<exists>x y. (P x y))" for P | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4987 | by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4988 | show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4989 | unfolding connected_closedin_eq | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4990 | proof (rule *) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4991 | fix H1 H2 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4992 | assume H: "closedin (subtopology euclidean (S \<union> T)) H1 \<and> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4993 | closedin (subtopology euclidean (S \<union> T)) H2 \<and> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4994 |                H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4995 | then have clo: "closedin (subtopology euclidean S) (S \<inter> H1)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4996 | "closedin (subtopology euclidean S) (S \<inter> H2)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4997 | by (metis Un_upper1 closedin_closed_subset inf_commute)+ | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4998 | have Seq: "S \<inter> (H1 \<union> H2) = S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 4999 | by (simp add: H) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5000 | have "S \<inter> ((S \<union> T) \<inter> H1) \<union> S \<inter> ((S \<union> T) \<inter> H2) = S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5001 | using Seq by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5002 |     moreover have "H1 \<inter> (S \<inter> ((S \<union> T) \<inter> H2)) = {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5003 | using H by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5004 |     ultimately have "S \<inter> H1 = {} \<or> S \<inter> H2 = {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5005 | by (metis (no_types) H Int_assoc \<open>S \<inter> (H1 \<union> H2) = S\<close> \<open>connected S\<close> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5006 | clo Seq connected_closedin inf_bot_right inf_le1) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5007 | then show "S \<subseteq> H1 \<or> S \<subseteq> H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5008 | using H \<open>connected S\<close> unfolding connected_closedin by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5009 | next | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5010 | fix H1 H2 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5011 | assume H: "closedin (subtopology euclidean (S \<union> T)) H1 \<and> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5012 | closedin (subtopology euclidean (S \<union> T)) H2 \<and> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5013 |                H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}" 
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5014 | and "S \<subseteq> H1" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5015 | then have H2T: "H2 \<subseteq> T" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5016 | by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5017 | have "T \<subseteq> U" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5018 | using Diff_iff opeT openin_imp_subset by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5019 | with \<open>S \<subseteq> U\<close> have Ueq: "U = (U - S) \<union> (S \<union> T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5020 | by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5021 | have "openin (subtopology euclidean ((U - S) \<union> (S \<union> T))) H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5022 | proof (rule openin_subtopology_Un) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5023 | show "openin (subtopology euclidean (S \<union> T)) H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5024 | using \<open>H2 \<subseteq> T\<close> apply (auto simp: openin_closedin_eq) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5025 | by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5026 | then show "openin (subtopology euclidean (U - S)) H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5027 | by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5028 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5029 | moreover have "closedin (subtopology euclidean ((U - S) \<union> (S \<union> T))) H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5030 | proof (rule closedin_subtopology_Un) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5031 | show "closedin (subtopology euclidean (U - S)) H2" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5032 | using H H2T cloT closedin_subset_trans | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5033 | by (blast intro: closedin_subtopology_Un closedin_trans) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5034 | qed (simp add: H) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5035 | ultimately | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5036 |     have H2: "H2 = {} \<or> H2 = U"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5037 | using Ueq \<open>connected U\<close> unfolding connected_clopen by metis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5038 | then have "H2 \<subseteq> S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5039 | by (metis Diff_partition H Un_Diff_cancel Un_subset_iff \<open>H2 \<subseteq> T\<close> assms(3) inf.orderE opeT openin_imp_subset) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5040 | moreover have "T \<subseteq> H2 - S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5041 | by (metis (no_types) H2 H opeT openin_closedin_eq topspace_euclidean_subtopology) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5042 | ultimately show False | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5043 | using H \<open>S \<subseteq> H1\<close> by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5044 | qed blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5045 | qed | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5046 | |
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5047 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 5048 | proposition component_diff_connected: | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5049 | fixes S :: "'a::metric_space set" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5050 | assumes "connected S" "connected U" "S \<subseteq> U" and C: "C \<in> components (U - S)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5051 | shows "connected(U - C)" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 5052 | using \<open>connected S\<close> unfolding connected_closedin_eq not_ex de_Morgan_conj | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5053 | proof clarify | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5054 | fix H3 H4 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5055 | assume clo3: "closedin (subtopology euclidean (U - C)) H3" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5056 | and clo4: "closedin (subtopology euclidean (U - C)) H4" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5057 |     and "H3 \<union> H4 = U - C" and "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5058 | and * [rule_format]: | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5059 | "\<forall>H1 H2. \<not> closedin (subtopology euclidean S) H1 \<or> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5060 | \<not> closedin (subtopology euclidean S) H2 \<or> | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5061 |                       H1 \<union> H2 \<noteq> S \<or> H1 \<inter> H2 \<noteq> {} \<or> \<not> H1 \<noteq> {} \<or> \<not> H2 \<noteq> {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5062 | then have "H3 \<subseteq> U-C" and ope3: "openin (subtopology euclidean (U - C)) (U - C - H3)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5063 | and "H4 \<subseteq> U-C" and ope4: "openin (subtopology euclidean (U - C)) (U - C - H4)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5064 | by (auto simp: closedin_def) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5065 |   have "C \<noteq> {}" "C \<subseteq> U-S" "connected C"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5066 | using C in_components_nonempty in_components_subset in_components_maximal by blast+ | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5067 | have cCH3: "connected (C \<union> H3)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5068 | proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo3]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5069 | show "openin (subtopology euclidean (U - C)) H3" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5070 | apply (simp add: openin_closedin_eq \<open>H3 \<subseteq> U - C\<close>) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5071 | apply (simp add: closedin_subtopology) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5072 |       by (metis Diff_cancel Diff_triv Un_Diff clo4 \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> closedin_closed inf_commute sup_bot.left_neutral)
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5073 | qed (use clo3 \<open>C \<subseteq> U - S\<close> in auto) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5074 | have cCH4: "connected (C \<union> H4)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5075 | proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo4]) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5076 | show "openin (subtopology euclidean (U - C)) H4" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5077 | apply (simp add: openin_closedin_eq \<open>H4 \<subseteq> U - C\<close>) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5078 | apply (simp add: closedin_subtopology) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5079 |       by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> clo3 closedin_closed)
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5080 | qed (use clo4 \<open>C \<subseteq> U - S\<close> in auto) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5081 | have "closedin (subtopology euclidean S) (S \<inter> H3)" "closedin (subtopology euclidean S) (S \<inter> H4)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5082 | using clo3 clo4 \<open>S \<subseteq> U\<close> \<open>C \<subseteq> U - S\<close> by (auto simp: closedin_closed) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5083 |   moreover have "S \<inter> H3 \<noteq> {}"      
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5084 |     using components_maximal [OF C cCH3] \<open>C \<noteq> {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H3 \<noteq> {}\<close> \<open>H3 \<subseteq> U - C\<close> by auto
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5085 |   moreover have "S \<inter> H4 \<noteq> {}"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5086 |     using components_maximal [OF C cCH4] \<open>C \<noteq> {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H4 \<noteq> {}\<close> \<open>H4 \<subseteq> U - C\<close> by auto
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5087 | ultimately show False | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5088 |     using * [of "S \<inter> H3" "S \<inter> H4"] \<open>H3 \<inter> H4 = {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H3 \<union> H4 = U - C\<close> \<open>S \<subseteq> U\<close> 
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5089 | by auto | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5090 | qed | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5091 | |
| 67962 | 5092 | subsection%unimportant\<open> Finite intersection property\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5093 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5094 | text\<open>Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5095 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5096 | lemma closed_imp_fip: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5097 | fixes S :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5098 | assumes "closed S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5099 | and T: "T \<in> \<F>" "bounded T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5100 | and clof: "\<And>T. T \<in> \<F> \<Longrightarrow> closed T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5101 |       and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5102 |     shows "S \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5103 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5104 | have "compact (S \<inter> T)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5105 | using \<open>closed S\<close> clof compact_eq_bounded_closed T by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5106 |   then have "(S \<inter> T) \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5107 | apply (rule compact_imp_fip) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5108 | apply (simp add: clof) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5109 | by (metis Int_assoc complete_lattice_class.Inf_insert finite_insert insert_subset none \<open>T \<in> \<F>\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5110 | then show ?thesis by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5111 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5112 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5113 | lemma closed_imp_fip_compact: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5114 | fixes S :: "'a::heine_borel set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5115 | shows | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5116 | "\<lbrakk>closed S; \<And>T. T \<in> \<F> \<Longrightarrow> compact T; | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5117 |      \<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> S \<inter> \<Inter>\<F>' \<noteq> {}\<rbrakk>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5118 |         \<Longrightarrow> S \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5119 | by (metis Inf_greatest closed_imp_fip compact_eq_bounded_closed empty_subsetI finite.emptyI inf.orderE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5120 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5121 | lemma closed_fip_heine_borel: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5122 | fixes \<F> :: "'a::heine_borel set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5123 | assumes "closed S" "T \<in> \<F>" "bounded T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5124 | and "\<And>T. T \<in> \<F> \<Longrightarrow> closed T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5125 |       and "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5126 |     shows "\<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5127 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5128 |   have "UNIV \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5129 | using assms closed_imp_fip [OF closed_UNIV] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5130 | then show ?thesis by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5131 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5132 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5133 | lemma compact_fip_heine_borel: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5134 | fixes \<F> :: "'a::heine_borel set set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5135 | assumes clof: "\<And>T. T \<in> \<F> \<Longrightarrow> compact T" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5136 |       and none: "\<And>\<F>'. \<lbrakk>finite \<F>'; \<F>' \<subseteq> \<F>\<rbrakk> \<Longrightarrow> \<Inter>\<F>' \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5137 |     shows "\<Inter>\<F> \<noteq> {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5138 | by (metis InterI all_not_in_conv clof closed_fip_heine_borel compact_eq_bounded_closed none) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5139 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5140 | lemma compact_sequence_with_limit: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5141 | fixes f :: "nat \<Rightarrow> 'a::heine_borel" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5142 | shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> compact (insert l (range f))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5143 | apply (simp add: compact_eq_bounded_closed, auto) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5144 | apply (simp add: convergent_imp_bounded) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5145 | by (simp add: closed_limpt islimpt_insert sequence_unique_limpt) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5146 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5147 | |
| 67962 | 5148 | subsection%unimportant\<open>Componentwise limits and continuity\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5149 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5150 | text\<open>But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\<close>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5151 | lemma Euclidean_dist_upper: "i \<in> Basis \<Longrightarrow> dist (x \<bullet> i) (y \<bullet> i) \<le> dist x y" | 
| 67155 | 5152 | by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5153 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5154 | text\<open>But is the premise @{term \<open>i \<in> Basis\<close>} really necessary?\<close>
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5155 | lemma open_preimage_inner: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5156 | assumes "open S" "i \<in> Basis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5157 |     shows "open {x. x \<bullet> i \<in> S}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5158 | proof (rule openI, simp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5159 | fix x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5160 | assume x: "x \<bullet> i \<in> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5161 | with assms obtain e where "0 < e" and e: "ball (x \<bullet> i) e \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5162 | by (auto simp: open_contains_ball_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5163 | have "\<exists>e>0. ball (y \<bullet> i) e \<subseteq> S" if dxy: "dist x y < e / 2" for y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5164 | proof (intro exI conjI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5165 | have "dist (x \<bullet> i) (y \<bullet> i) < e / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5166 | by (meson \<open>i \<in> Basis\<close> dual_order.trans Euclidean_dist_upper not_le that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5167 | then have "dist (x \<bullet> i) z < e" if "dist (y \<bullet> i) z < e / 2" for z | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5168 | by (metis dist_commute dist_triangle_half_l that) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5169 | then have "ball (y \<bullet> i) (e / 2) \<subseteq> ball (x \<bullet> i) e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5170 | using mem_ball by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5171 | with e show "ball (y \<bullet> i) (e / 2) \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5172 | by (metis order_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5173 | qed (simp add: \<open>0 < e\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5174 |   then show "\<exists>e>0. ball x e \<subseteq> {s. s \<bullet> i \<in> S}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5175 | by (metis (no_types, lifting) \<open>0 < e\<close> \<open>open S\<close> half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5176 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5177 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 5178 | proposition tendsto_componentwise_iff: | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5179 | fixes f :: "_ \<Rightarrow> 'b::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5180 | shows "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>i \<in> Basis. ((\<lambda>x. (f x \<bullet> i)) \<longlongrightarrow> (l \<bullet> i)) F)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5181 | (is "?lhs = ?rhs") | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68527diff
changeset | 5182 | proof | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5183 | assume ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5184 | then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5185 | unfolding tendsto_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5186 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5187 |     apply (drule_tac x="{s. s \<bullet> i \<in> S}" in spec)
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5188 | apply (auto simp: open_preimage_inner) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5189 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5190 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5191 | assume R: ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5192 | then have "\<And>e. e > 0 \<Longrightarrow> \<forall>i\<in>Basis. \<forall>\<^sub>F x in F. dist (f x \<bullet> i) (l \<bullet> i) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5193 | unfolding tendsto_iff by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5194 | then have R': "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5195 | by (simp add: eventually_ball_finite_distrib [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5196 | show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5197 | unfolding tendsto_iff | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5198 | proof clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5199 | fix e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5200 | assume "0 < e" | 
| 67155 | 5201 | have *: "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5202 |              if "\<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / real DIM('b)" for x
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5203 | proof - | 
| 67155 | 5204 | have "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> sum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis" | 
| 5205 | by (simp add: L2_set_le_sum) | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5206 |       also have "... < DIM('b) * (e / real DIM('b))"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5207 | apply (rule sum_bounded_above_strict) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5208 | using that by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5209 | also have "... = e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5210 | by (simp add: field_simps) | 
| 67155 | 5211 | finally show "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" . | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5212 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5213 |     have "\<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / DIM('b)"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5214 | apply (rule R') | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5215 | using \<open>0 < e\<close> by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5216 | then show "\<forall>\<^sub>F x in F. dist (f x) l < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5217 | apply (rule eventually_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5218 | apply (subst euclidean_dist_l2) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5219 | using * by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5220 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5221 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5222 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5223 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5224 | corollary continuous_componentwise: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5225 | "continuous F f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous F (\<lambda>x. (f x \<bullet> i)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5226 | by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5227 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5228 | corollary continuous_on_componentwise: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5229 | fixes S :: "'a :: t2_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5230 | shows "continuous_on S f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous_on S (\<lambda>x. (f x \<bullet> i)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5231 | apply (simp add: continuous_on_eq_continuous_within) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5232 | using continuous_componentwise by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5233 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5234 | lemma linear_componentwise_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5235 | "(linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. linear (\<lambda>x. f' x \<bullet> i))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5236 | apply (auto simp: linear_iff inner_left_distrib) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5237 | apply (metis inner_left_distrib euclidean_eq_iff) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5238 | by (metis euclidean_eqI inner_scaleR_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5239 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5240 | lemma bounded_linear_componentwise_iff: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5241 | "(bounded_linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. bounded_linear (\<lambda>x. f' x \<bullet> i))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5242 | (is "?lhs = ?rhs") | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5243 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5244 | assume ?lhs then show ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5245 | by (simp add: bounded_linear_inner_left_comp) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5246 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5247 | assume ?rhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5248 | then have "(\<forall>i\<in>Basis. \<exists>K. \<forall>x. \<bar>f' x \<bullet> i\<bar> \<le> norm x * K)" "linear f'" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5249 | by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5250 | then obtain F where F: "\<And>i x. i \<in> Basis \<Longrightarrow> \<bar>f' x \<bullet> i\<bar> \<le> norm x * F i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5251 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5252 | have "norm (f' x) \<le> norm x * sum F Basis" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5253 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5254 | have "norm (f' x) \<le> (\<Sum>i\<in>Basis. \<bar>f' x \<bullet> i\<bar>)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5255 | by (rule norm_le_l1) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5256 | also have "... \<le> (\<Sum>i\<in>Basis. norm x * F i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5257 | by (metis F sum_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5258 | also have "... = norm x * sum F Basis" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5259 | by (simp add: sum_distrib_left) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5260 | finally show ?thesis . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5261 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5262 | then show ?lhs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5263 | by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5264 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5265 | |
| 67962 | 5266 | subsection%unimportant\<open>Pasting functions together\<close> | 
| 5267 | ||
| 5268 | subsubsection%unimportant\<open>on open sets\<close> | |
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5269 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5270 | lemma pasting_lemma: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5271 | fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5272 | assumes clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5273 | and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5274 | and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5275 | and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5276 | shows "continuous_on S g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5277 | proof (clarsimp simp: continuous_openin_preimage_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5278 | fix U :: "'b set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5279 | assume "open U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5280 | have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5281 | using clo openin_imp_subset by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5282 | have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5283 | using S f g by fastforce | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5284 | show "openin (subtopology euclidean S) (S \<inter> g -` U)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5285 | apply (subst *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5286 | apply (rule openin_Union, clarify) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5287 | using \<open>open U\<close> clo cont continuous_openin_preimage_gen openin_trans by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5288 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5289 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5290 | lemma pasting_lemma_exists: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5291 | fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5292 | assumes S: "S \<subseteq> (\<Union>i \<in> I. T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5293 | and clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5294 | and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5295 | and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5296 | obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5297 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5298 | show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5299 | apply (rule pasting_lemma [OF clo cont]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5300 | apply (blast intro: f)+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5301 | apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5302 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5303 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5304 | fix x i | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5305 | assume "i \<in> I" "x \<in> S \<inter> T i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5306 | then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5307 | by (metis (no_types, lifting) IntD2 IntI f someI_ex) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5308 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5309 | |
| 67962 | 5310 | subsubsection%unimportant\<open>Likewise on closed sets, with a finiteness assumption\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5311 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5312 | lemma pasting_lemma_closed: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5313 | fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5314 | assumes "finite I" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5315 | and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5316 | and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5317 | and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5318 | and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5319 | shows "continuous_on S g" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5320 | proof (clarsimp simp: continuous_closedin_preimage_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5321 | fix U :: "'b set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5322 | assume "closed U" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5323 | have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5324 | using clo closedin_imp_subset by blast | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5325 | have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5326 | using S f g by fastforce | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5327 | show "closedin (subtopology euclidean S) (S \<inter> g -` U)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5328 | apply (subst *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5329 | apply (rule closedin_Union) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5330 | using \<open>finite I\<close> apply simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5331 | apply (blast intro: \<open>closed U\<close> continuous_closedin_preimage cont clo closedin_trans) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5332 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5333 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5334 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5335 | lemma pasting_lemma_exists_closed: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5336 | fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5337 | assumes "finite I" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5338 | and S: "S \<subseteq> (\<Union>i \<in> I. T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5339 | and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5340 | and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5341 | and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5342 | obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5343 | proof | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5344 | show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5345 | apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5346 | apply (blast intro: f)+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5347 | apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5348 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5349 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5350 | fix x i | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5351 | assume "i \<in> I" "x \<in> S \<inter> T i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5352 | then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5353 | by (metis (no_types, lifting) IntD2 IntI f someI_ex) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5354 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5355 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5356 | lemma tube_lemma: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5357 | assumes "compact K" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5358 | assumes "open W" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5359 |   assumes "{x0} \<times> K \<subseteq> W"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5360 | shows "\<exists>X0. x0 \<in> X0 \<and> open X0 \<and> X0 \<times> K \<subseteq> W" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5361 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5362 |   {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5363 | fix y assume "y \<in> K" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5364 | then have "(x0, y) \<in> W" using assms by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5365 | with \<open>open W\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5366 | have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5367 | by (rule open_prod_elim) blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5368 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5369 | then obtain X0 Y where | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5370 | *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5371 | by metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5372 | from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5373 | with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5374 | by (meson compactE) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5375 | then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5376 | by (force intro!: choice) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5377 | with * CC show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5378 | by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"]) (* SLOW *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5379 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5380 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5381 | lemma continuous_on_prod_compactE: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5382 | fixes fx::"'a::topological_space \<times> 'b::topological_space \<Rightarrow> 'c::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5383 | and e::real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5384 | assumes cont_fx: "continuous_on (U \<times> C) fx" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5385 | assumes "compact C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5386 | assumes [intro]: "x0 \<in> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5387 | notes [continuous_intros] = continuous_on_compose2[OF cont_fx] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5388 | assumes "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5389 | obtains X0 where "x0 \<in> X0" "open X0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5390 | "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5391 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5392 | define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5393 |   define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5394 |   have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5395 | by (auto simp: vimage_def W0_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5396 |   have "open {..<e}" by simp
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5397 | have "continuous_on (U \<times> C) psi" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5398 | by (auto intro!: continuous_intros simp: psi_def split_beta') | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5399 |   from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5400 | obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5401 | unfolding W0_eq by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5402 |   have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5403 | unfolding W | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5404 | by (auto simp: W0_def psi_def \<open>0 < e\<close>) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5405 |   then have "{x0} \<times> C \<subseteq> W" by blast
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5406 | from tube_lemma[OF \<open>compact C\<close> \<open>open W\<close> this] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5407 | obtain X0 where X0: "x0 \<in> X0" "open X0" "X0 \<times> C \<subseteq> W" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5408 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5409 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5410 | have "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5411 | proof safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5412 | fix x assume x: "x \<in> X0" "x \<in> U" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5413 | fix t assume t: "t \<in> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5414 | have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5415 | by (auto simp: psi_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5416 | also | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5417 |     {
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5418 | have "(x, t) \<in> X0 \<times> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5419 | using t x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5420 | by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5421 | also note \<open>\<dots> \<subseteq> W\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5422 | finally have "(x, t) \<in> W" . | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5423 | with t x have "(x, t) \<in> W \<inter> U \<times> C" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5424 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5425 | also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5426 | finally have "psi (x, t) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5427 | by (auto simp: W0_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5428 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5429 | finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5430 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5431 | from X0(1,2) this show ?thesis .. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5432 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5433 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5434 | |
| 67962 | 5435 | subsection%unimportant\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5436 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5437 | text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5438 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5439 | lemma continuous_disconnected_range_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5440 | assumes S: "connected S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5441 | and conf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5442 | and fim: "f ` S \<subseteq> t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5443 |       and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5444 | shows "f constant_on S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5445 | proof (cases "S = {}")
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5446 | case True then show ?thesis | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5447 | by (simp add: constant_on_def) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5448 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5449 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5450 |   { fix x assume "x \<in> S"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5451 |     then have "f ` S \<subseteq> {f x}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5452 | by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5453 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5454 | with False show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5455 | unfolding constant_on_def by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5456 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5457 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5458 | lemma discrete_subset_disconnected: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5459 | fixes S :: "'a::topological_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5460 | fixes t :: "'b::real_normed_vector set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5461 | assumes conf: "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5462 | and no: "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5463 |    shows "f ` S \<subseteq> {y. connected_component_set (f ` S) y = {y}}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5464 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5465 |   { fix x assume x: "x \<in> S"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5466 | then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> S; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5467 | using conf no [OF x] by auto | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5468 | then have e2: "0 \<le> e / 2" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5469 | by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5470 | have "f y = f x" if "y \<in> S" and ccs: "f y \<in> connected_component_set (f ` S) (f x)" for y | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5471 | apply (rule ccontr) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5472 | using connected_closed [of "connected_component_set (f ` S) (f x)"] \<open>e>0\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5473 | apply (simp add: del: ex_simps) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5474 | apply (drule spec [where x="cball (f x) (e / 2)"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5475 | apply (drule spec [where x="- ball(f x) e"]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5476 | apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5477 | apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5478 | using centre_in_cball connected_component_refl_eq e2 x apply blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5479 | using ccs | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5480 | apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> S\<close>]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5481 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5482 | moreover have "connected_component_set (f ` S) (f x) \<subseteq> f ` S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5483 | by (auto simp: connected_component_in) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5484 |     ultimately have "connected_component_set (f ` S) (f x) = {f x}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5485 | by (auto simp: x) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5486 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5487 | with assms show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5488 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5489 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5490 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5491 | lemma finite_implies_discrete: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5492 | fixes S :: "'a::topological_space set" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5493 | assumes "finite (f ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5494 | shows "(\<forall>x \<in> S. \<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5495 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5496 | have "\<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> S" for x | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5497 |   proof (cases "f ` S - {f x} = {}")
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5498 | case True | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5499 | with zero_less_numeral show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5500 | by (fastforce simp add: Set.image_subset_iff cong: conj_cong) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5501 | next | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5502 | case False | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5503 | then obtain z where z: "z \<in> S" "f z \<noteq> f x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5504 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5505 |     have finn: "finite {norm (z - f x) |z. z \<in> f ` S - {f x}}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5506 | using assms by simp | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5507 |     then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` S - {f x}}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5508 | apply (rule finite_imp_less_Inf) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5509 | using z apply force+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5510 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5511 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5512 | by (force intro!: * cInf_le_finite [OF finn]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5513 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5514 | with assms show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5515 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5516 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5517 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5518 | text\<open>This proof requires the existence of two separate values of the range type.\<close> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5519 | lemma finite_range_constant_imp_connected: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5520 | assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5521 | \<lbrakk>continuous_on S f; finite(f ` S)\<rbrakk> \<Longrightarrow> f constant_on S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5522 | shows "connected S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5523 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5524 |   { fix t u
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5525 | assume clt: "closedin (subtopology euclidean S) t" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5526 | and clu: "closedin (subtopology euclidean S) u" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5527 |        and tue: "t \<inter> u = {}" and tus: "t \<union> u = S"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5528 | have conif: "continuous_on S (\<lambda>x. if x \<in> t then 0 else 1)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5529 | apply (subst tus [symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5530 | apply (rule continuous_on_cases_local) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5531 | using clt clu tue | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5532 | apply (auto simp: tus continuous_on_const) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5533 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5534 | have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` S)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5535 |       by (rule finite_subset [of _ "{0,1}"]) auto
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5536 |     have "t = {} \<or> u = {}"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5537 | using assms [OF conif fi] tus [symmetric] | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5538 | by (auto simp: Ball_def constant_on_def) (metis IntI empty_iff one_neq_zero tue) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5539 | } | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5540 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5541 | by (simp add: connected_closedin_eq) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5542 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5543 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5544 | lemma continuous_disconnected_range_constant_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5545 | "(connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5546 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5547 |             \<forall>t. continuous_on S f \<and> f ` S \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5548 | \<longrightarrow> f constant_on S))" (is ?thesis1) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5549 | and continuous_discrete_range_constant_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5550 | "(connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5551 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5552 | continuous_on S f \<and> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5553 | (\<forall>x \<in> S. \<exists>e. 0 < e \<and> (\<forall>y. y \<in> S \<and> (f y \<noteq> f x) \<longrightarrow> e \<le> norm(f y - f x))) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5554 | \<longrightarrow> f constant_on S))" (is ?thesis2) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5555 | and continuous_finite_range_constant_eq: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5556 | "(connected S \<longleftrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5557 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5558 | continuous_on S f \<and> finite (f ` S) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5559 | \<longrightarrow> f constant_on S))" (is ?thesis3) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5560 | proof - | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5561 | have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5562 | \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5563 | by blast | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5564 | have "?thesis1 \<and> ?thesis2 \<and> ?thesis3" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5565 | apply (rule *) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5566 | using continuous_disconnected_range_constant apply metis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5567 | apply clarify | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5568 | apply (frule discrete_subset_disconnected; blast) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5569 | apply (blast dest: finite_implies_discrete) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5570 | apply (blast intro!: finite_range_constant_imp_connected) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5571 | done | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5572 | then show ?thesis1 ?thesis2 ?thesis3 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5573 | by blast+ | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5574 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5575 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5576 | lemma continuous_discrete_range_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5577 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5578 | assumes S: "connected S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5579 | and "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5580 | and "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5581 | shows "f constant_on S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5582 | using continuous_discrete_range_constant_eq [THEN iffD1, OF S] assms by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5583 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5584 | lemma continuous_finite_range_constant: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5585 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5586 | assumes "connected S" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5587 | and "continuous_on S f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5588 | and "finite (f ` S)" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5589 | shows "f constant_on S" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66835diff
changeset | 5590 | using assms continuous_finite_range_constant_eq by blast | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5591 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5592 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5593 | |
| 67962 | 5594 | subsection%unimportant \<open>Continuous Extension\<close> | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5595 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5596 | definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5597 | "clamp a b x = (if (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5598 | then (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5599 | else a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5600 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5601 | lemma clamp_in_interval[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5602 | assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5603 | shows "clamp a b x \<in> cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5604 | unfolding clamp_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5605 | using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5606 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5607 | lemma clamp_cancel_cbox[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5608 | fixes x a b :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5609 | assumes x: "x \<in> cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5610 | shows "clamp a b x = x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5611 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5612 | by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5613 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5614 | lemma clamp_empty_interval: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5615 | assumes "i \<in> Basis" "a \<bullet> i > b \<bullet> i" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5616 | shows "clamp a b = (\<lambda>_. a)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5617 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5618 | by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5619 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5620 | lemma dist_clamps_le_dist_args: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5621 | fixes x :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5622 | shows "dist (clamp a b y) (clamp a b x) \<le> dist y x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5623 | proof cases | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5624 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5625 | then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5626 | (\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5627 | by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5628 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5629 | by (auto intro: real_sqrt_le_mono | 
| 67155 | 5630 | simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5631 | qed (auto simp: clamp_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5632 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5633 | lemma clamp_continuous_at: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5634 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5635 | and x :: 'a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5636 | assumes f_cont: "continuous_on (cbox a b) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5637 | shows "continuous (at x) (\<lambda>x. f (clamp a b x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5638 | proof cases | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5639 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5640 | show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5641 | unfolding continuous_at_eps_delta | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5642 | proof safe | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5643 | fix x :: 'a | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5644 | fix e :: real | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5645 | assume "e > 0" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5646 | moreover have "clamp a b x \<in> cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5647 | by (simp add: clamp_in_interval le) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5648 | moreover note f_cont[simplified continuous_on_iff] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5649 | ultimately | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5650 | obtain d where d: "0 < d" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5651 | "\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5652 | by force | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5653 | show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5654 | dist (f (clamp a b x')) (f (clamp a b x)) < e" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5655 | using le | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5656 | by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5657 | qed | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5658 | qed (auto simp: clamp_empty_interval) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5659 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5660 | lemma clamp_continuous_on: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5661 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5662 | assumes f_cont: "continuous_on (cbox a b) f" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5663 | shows "continuous_on S (\<lambda>x. f (clamp a b x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5664 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5665 | by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5666 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5667 | lemma clamp_bounded: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5668 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5669 | assumes bounded: "bounded (f ` (cbox a b))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5670 | shows "bounded (range (\<lambda>x. f (clamp a b x)))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5671 | proof cases | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5672 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5673 | from bounded obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. dist undefined x \<le> c" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5674 | by (auto simp: bounded_any_center[where a=undefined]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5675 | then show ?thesis | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5676 | by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5677 | simp: bounded_any_center[where a=undefined]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5678 | qed (auto simp: clamp_empty_interval image_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5679 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5680 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5681 | definition ext_cont :: "('a::euclidean_space \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5682 | where "ext_cont f a b = (\<lambda>x. f (clamp a b x))" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5683 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5684 | lemma ext_cont_cancel_cbox[simp]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5685 | fixes x a b :: "'a::euclidean_space" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5686 | assumes x: "x \<in> cbox a b" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5687 | shows "ext_cont f a b x = f x" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5688 | using assms | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5689 | unfolding ext_cont_def | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5690 | by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5691 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5692 | lemma continuous_on_ext_cont[continuous_intros]: | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5693 | "continuous_on (cbox a b) f \<Longrightarrow> continuous_on S (ext_cont f a b)" | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5694 | by (auto intro!: clamp_continuous_on simp: ext_cont_def) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5695 | |
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5696 | end |