| author | wenzelm | 
| Thu, 04 Nov 2021 12:25:23 +0100 | |
| changeset 74683 | c8327efc7af1 | 
| parent 73791 | e10d530f157a | 
| child 76724 | 7ff71bdcf731 | 
| permissions | -rw-r--r-- | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1 | (* Author: L C Paulson, University of Cambridge | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 2 | Author: Amine Chaieb, University of Cambridge | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 3 | Author: Robert Himmelmann, TU Muenchen | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 4 | Author: Brian Huffman, Portland State University | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 5 | *) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 6 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 7 | section \<open>Elementary Normed Vector Spaces\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 8 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 9 | theory Elementary_Normed_Spaces | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 10 | imports | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 11 | "HOL-Library.FuncSet" | 
| 70724 | 12 | Elementary_Metric_Spaces Cartesian_Space | 
| 69617 | 13 | Connected | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 14 | begin | 
| 70724 | 15 | subsection \<open>Orthogonal Transformation of Balls\<close> | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 16 | |
| 70136 | 17 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas Combining Imports\<close> | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 18 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 19 | lemma open_sums: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 20 |   fixes T :: "('b::real_normed_vector) set"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 21 | assumes "open S \<or> open T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 22 |   shows "open (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 23 | using assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 24 | proof | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 25 | assume S: "open S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 26 | show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 27 | proof (clarsimp simp: open_dist) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 28 | fix x y | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 29 | assume "x \<in> S" "y \<in> T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 30 | with S obtain e where "e > 0" and e: "\<And>x'. dist x' x < e \<Longrightarrow> x' \<in> S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 31 | by (auto simp: open_dist) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 32 | then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 33 | by (metis \<open>y \<in> T\<close> diff_add_cancel dist_add_cancel2) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 34 | then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 35 | using \<open>0 < e\<close> \<open>x \<in> S\<close> by blast | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 36 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 37 | next | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 38 | assume T: "open T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 39 | show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 40 | proof (clarsimp simp: open_dist) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 41 | fix x y | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 42 | assume "x \<in> S" "y \<in> T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 43 | with T obtain e where "e > 0" and e: "\<And>x'. dist x' y < e \<Longrightarrow> x' \<in> T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 44 | by (auto simp: open_dist) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 45 | then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 46 | by (metis \<open>x \<in> S\<close> add_diff_cancel_left' add_diff_eq diff_diff_add dist_norm) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 47 | then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 48 | using \<open>0 < e\<close> \<open>y \<in> T\<close> by blast | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 49 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 50 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 51 | |
| 70724 | 52 | lemma image_orthogonal_transformation_ball: | 
| 53 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | |
| 54 | assumes "orthogonal_transformation f" | |
| 55 | shows "f ` ball x r = ball (f x) r" | |
| 56 | proof (intro equalityI subsetI) | |
| 57 | fix y assume "y \<in> f ` ball x r" | |
| 58 | with assms show "y \<in> ball (f x) r" | |
| 59 | by (auto simp: orthogonal_transformation_isometry) | |
| 60 | next | |
| 61 | fix y assume y: "y \<in> ball (f x) r" | |
| 62 | then obtain z where z: "y = f z" | |
| 63 | using assms orthogonal_transformation_surj by blast | |
| 64 | with y assms show "y \<in> f ` ball x r" | |
| 65 | by (auto simp: orthogonal_transformation_isometry) | |
| 66 | qed | |
| 67 | ||
| 68 | lemma image_orthogonal_transformation_cball: | |
| 69 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | |
| 70 | assumes "orthogonal_transformation f" | |
| 71 | shows "f ` cball x r = cball (f x) r" | |
| 72 | proof (intro equalityI subsetI) | |
| 73 | fix y assume "y \<in> f ` cball x r" | |
| 74 | with assms show "y \<in> cball (f x) r" | |
| 75 | by (auto simp: orthogonal_transformation_isometry) | |
| 76 | next | |
| 77 | fix y assume y: "y \<in> cball (f x) r" | |
| 78 | then obtain z where z: "y = f z" | |
| 79 | using assms orthogonal_transformation_surj by blast | |
| 80 | with y assms show "y \<in> f ` cball x r" | |
| 81 | by (auto simp: orthogonal_transformation_isometry) | |
| 82 | qed | |
| 83 | ||
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 84 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 85 | subsection \<open>Support\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 86 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 87 | definition (in monoid_add) support_on :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'b set"
 | 
| 72643 | 88 |   where "support_on S f = {x\<in>S. f x \<noteq> 0}"
 | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 89 | |
| 72643 | 90 | lemma in_support_on: "x \<in> support_on S f \<longleftrightarrow> x \<in> S \<and> f x \<noteq> 0" | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 91 | by (simp add: support_on_def) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 92 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 93 | lemma support_on_simps[simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 94 |   "support_on {} f = {}"
 | 
| 72643 | 95 | "support_on (insert x S) f = | 
| 96 | (if f x = 0 then support_on S f else insert x (support_on S f))" | |
| 97 | "support_on (S \<union> T) f = support_on S f \<union> support_on T f" | |
| 98 | "support_on (S \<inter> T) f = support_on S f \<inter> support_on T f" | |
| 99 | "support_on (S - T) f = support_on S f - support_on T f" | |
| 100 | "support_on (f ` S) g = f ` (support_on S (g \<circ> f))" | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 101 | unfolding support_on_def by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 102 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 103 | lemma support_on_cong: | 
| 72643 | 104 | "(\<And>x. x \<in> S \<Longrightarrow> f x = 0 \<longleftrightarrow> g x = 0) \<Longrightarrow> support_on S f = support_on S g" | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 105 | by (auto simp: support_on_def) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 106 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 107 | lemma support_on_if: "a \<noteq> 0 \<Longrightarrow> support_on A (\<lambda>x. if P x then a else 0) = {x\<in>A. P x}"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 108 | by (auto simp: support_on_def) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 109 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 110 | lemma support_on_if_subset: "support_on A (\<lambda>x. if P x then a else 0) \<subseteq> {x \<in> A. P x}"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 111 | by (auto simp: support_on_def) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 112 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 113 | lemma finite_support[intro]: "finite S \<Longrightarrow> finite (support_on S f)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 114 | unfolding support_on_def by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 115 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 116 | (* TODO: is supp_sum really needed? TODO: Generalize to Finite_Set.fold *) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 117 | definition (in comm_monoid_add) supp_sum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 118 | where "supp_sum f S = (\<Sum>x\<in>support_on S f. f x)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 119 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 120 | lemma supp_sum_empty[simp]: "supp_sum f {} = 0"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 121 | unfolding supp_sum_def by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 122 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 123 | lemma supp_sum_insert[simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 124 | "finite (support_on S f) \<Longrightarrow> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 125 | supp_sum f (insert x S) = (if x \<in> S then supp_sum f S else f x + supp_sum f S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 126 | by (simp add: supp_sum_def in_support_on insert_absorb) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 127 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 128 | lemma supp_sum_divide_distrib: "supp_sum f A / (r::'a::field) = supp_sum (\<lambda>n. f n / r) A" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 129 | by (cases "r = 0") | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 130 | (auto simp: supp_sum_def sum_divide_distrib intro!: sum.cong support_on_cong) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 131 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 132 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 133 | subsection \<open>Intervals\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 134 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 135 | lemma image_affinity_interval: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 136 | fixes c :: "'a::ordered_real_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 137 |   shows "((\<lambda>x. m *\<^sub>R x + c) ` {a..b}) = 
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 138 |            (if {a..b}={} then {}
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 139 |             else if 0 \<le> m then {m *\<^sub>R a + c .. m  *\<^sub>R b + c}
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 140 |             else {m *\<^sub>R b + c .. m *\<^sub>R a + c})"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 141 | (is "?lhs = ?rhs") | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 142 | proof (cases "m=0") | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 143 | case True | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 144 | then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 145 | by force | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 146 | next | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 147 | case False | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 148 | show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 149 | proof | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 150 | show "?lhs \<subseteq> ?rhs" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 151 | by (auto simp: scaleR_left_mono scaleR_left_mono_neg) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 152 | show "?rhs \<subseteq> ?lhs" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 153 | proof (clarsimp, intro conjI impI subsetI) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 154 |       show "\<lbrakk>0 \<le> m; a \<le> b; x \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}\<rbrakk>
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 155 |             \<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
 | 
| 72643 | 156 | using False | 
| 157 | by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) | |
| 158 | (auto simp: pos_le_divideR_eq pos_divideR_le_eq le_diff_eq diff_le_eq) | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 159 |       show "\<lbrakk>\<not> 0 \<le> m; a \<le> b;  x \<in> {m *\<^sub>R b + c..m *\<^sub>R a + c}\<rbrakk>
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 160 |             \<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
 | 
| 72643 | 161 | by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) | 
| 162 | (auto simp add: neg_le_divideR_eq neg_divideR_le_eq le_diff_eq diff_le_eq) | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 163 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 164 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 165 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 166 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 167 | subsection \<open>Limit Points\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 168 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 169 | lemma islimpt_ball: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 170 |   fixes x y :: "'a::{real_normed_vector,perfect_space}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 171 | shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 172 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 173 | proof | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 174 | show ?rhs if ?lhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 175 | proof | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 176 |     {
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 177 | assume "e \<le> 0" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 178 |       then have *: "ball x e = {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 179 | using ball_eq_empty[of x e] by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 180 | have False using \<open>?lhs\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 181 | unfolding * using islimpt_EMPTY[of y] by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 182 | } | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 183 | then show "e > 0" by (metis not_less) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 184 | show "y \<in> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 185 | using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 186 | ball_subset_cball[of x e] \<open>?lhs\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 187 | unfolding closed_limpt by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 188 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 189 | show ?lhs if ?rhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 190 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 191 | from that have "e > 0" by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 192 |     {
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 193 | fix d :: real | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 194 | assume "d > 0" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 195 | have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 196 | proof (cases "d \<le> dist x y") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 197 | case True | 
| 72643 | 198 | then show ?thesis | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 199 | proof (cases "x = y") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 200 | case True | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 201 | then have False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 202 | using \<open>d \<le> dist x y\<close> \<open>d>0\<close> by auto | 
| 72643 | 203 | then show ?thesis | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 204 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 205 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 206 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 207 | have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) = | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 208 | norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 209 | unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 210 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 211 | also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 212 | using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 213 | unfolding scaleR_minus_left scaleR_one | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 214 | by (auto simp: norm_minus_commute) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 215 | also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 216 | unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 217 | unfolding distrib_right using \<open>x\<noteq>y\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 218 | also have "\<dots> \<le> e - d/2" using \<open>d \<le> dist x y\<close> and \<open>d>0\<close> and \<open>?rhs\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 219 | by (auto simp: dist_norm) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 220 | finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using \<open>d>0\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 221 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 222 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 223 | have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 224 | using \<open>x\<noteq>y\<close>[unfolded dist_nz] \<open>d>0\<close> unfolding scaleR_eq_0_iff | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 225 | by (auto simp: dist_commute) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 226 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 227 | have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" | 
| 72643 | 228 | using \<open>0 < d\<close> by (fastforce simp: dist_norm) | 
| 229 | ultimately show ?thesis | |
| 230 | by (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 231 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 232 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 233 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 234 | then have "d > dist x y" by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 235 | show "\<exists>x' \<in> ball x e. x' \<noteq> y \<and> dist x' y < d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 236 | proof (cases "x = y") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 237 | case True | 
| 72643 | 238 | obtain z where z: "z \<noteq> y" "dist z y < min e d" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 239 | using perfect_choose_dist[of "min e d" y] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 240 | using \<open>d > 0\<close> \<open>e>0\<close> by auto | 
| 72643 | 241 | show ?thesis | 
| 242 | by (metis True z dist_commute mem_ball min_less_iff_conj) | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 243 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 244 | case False | 
| 72643 | 245 | then show ?thesis | 
| 246 | using \<open>d>0\<close> \<open>d > dist x y\<close> \<open>?rhs\<close> by force | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 247 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 248 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 249 | } | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 250 | then show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 251 | unfolding mem_cball islimpt_approachable mem_ball by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 252 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 253 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 254 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 255 | lemma closure_ball_lemma: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 256 | fixes x y :: "'a::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 257 | assumes "x \<noteq> y" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 258 | shows "y islimpt ball x (dist x y)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 259 | proof (rule islimptI) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 260 | fix T | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 261 | assume "y \<in> T" "open T" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 262 | then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 263 | unfolding open_dist by fast | 
| 72643 | 264 |   \<comment>\<open>choose point between @{term x} and @{term y}, within distance @{term r} of @{term y}.\<close>
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 265 | define k where "k = min 1 (r / (2 * dist x y))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 266 | define z where "z = y + scaleR k (x - y)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 267 | have z_def2: "z = x + scaleR (1 - k) (y - x)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 268 | unfolding z_def by (simp add: algebra_simps) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 269 | have "dist z y < r" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 270 | unfolding z_def k_def using \<open>0 < r\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 271 | by (simp add: dist_norm min_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 272 | then have "z \<in> T" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 273 | using \<open>\<forall>z. dist z y < r \<longrightarrow> z \<in> T\<close> by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 274 | have "dist x z < dist x y" | 
| 72643 | 275 | using \<open>0 < r\<close> assms by (simp add: z_def2 k_def dist_norm norm_minus_commute) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 276 | then have "z \<in> ball x (dist x y)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 277 | by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 278 | have "z \<noteq> y" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 279 | unfolding z_def k_def using \<open>x \<noteq> y\<close> \<open>0 < r\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 280 | by (simp add: min_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 281 | show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 282 | using \<open>z \<in> ball x (dist x y)\<close> \<open>z \<in> T\<close> \<open>z \<noteq> y\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 283 | by fast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 284 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 285 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 286 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 287 | subsection \<open>Balls and Spheres in Normed Spaces\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 288 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 289 | lemma mem_ball_0 [simp]: "x \<in> ball 0 e \<longleftrightarrow> norm x < e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 290 | for x :: "'a::real_normed_vector" | 
| 72643 | 291 | by simp | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 292 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 293 | lemma mem_cball_0 [simp]: "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 294 | for x :: "'a::real_normed_vector" | 
| 72643 | 295 | by simp | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 296 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 297 | lemma closure_ball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 298 | fixes x :: "'a::real_normed_vector" | 
| 72643 | 299 | assumes "0 < e" | 
| 300 | shows "closure (ball x e) = cball x e" | |
| 301 | proof | |
| 302 | show "closure (ball x e) \<subseteq> cball x e" | |
| 303 | using closed_cball closure_minimal by blast | |
| 304 | have "\<And>y. dist x y < e \<or> dist x y = e \<Longrightarrow> y \<in> closure (ball x e)" | |
| 305 | by (metis Un_iff assms closure_ball_lemma closure_def dist_eq_0_iff mem_Collect_eq mem_ball) | |
| 306 | then show "cball x e \<subseteq> closure (ball x e)" | |
| 307 | by force | |
| 308 | qed | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 309 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 310 | lemma mem_sphere_0 [simp]: "x \<in> sphere 0 e \<longleftrightarrow> norm x = e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 311 | for x :: "'a::real_normed_vector" | 
| 72643 | 312 | by simp | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 313 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 314 | (* In a trivial vector space, this fails for e = 0. *) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 315 | lemma interior_cball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 316 |   fixes x :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 317 | shows "interior (cball x e) = ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 318 | proof (cases "e \<ge> 0") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 319 | case False note cs = this | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 320 |   from cs have null: "ball x e = {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 321 | using ball_empty[of e x] by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 322 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 323 |   have "cball x e = {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 324 | proof (rule equals0I) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 325 | fix y | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 326 | assume "y \<in> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 327 | then show False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 328 | by (metis ball_eq_empty null cs dist_eq_0_iff dist_le_zero_iff empty_subsetI mem_cball | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 329 | subset_antisym subset_ball) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 330 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 331 |   then have "interior (cball x e) = {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 332 | using interior_empty by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 333 | ultimately show ?thesis by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 334 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 335 | case True note cs = this | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 336 | have "ball x e \<subseteq> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 337 | using ball_subset_cball by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 338 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 339 |   {
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 340 | fix S y | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 341 | assume as: "S \<subseteq> cball x e" "open S" "y\<in>S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 342 | then obtain d where "d>0" and d: "\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 343 | unfolding open_dist by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 344 | then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 345 | using perfect_choose_dist [of d] by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 346 | have "xa \<in> S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 347 | using d[THEN spec[where x = xa]] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 348 | using xa by (auto simp: dist_commute) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 349 | then have xa_cball: "xa \<in> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 350 | using as(1) by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 351 | then have "y \<in> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 352 | proof (cases "x = y") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 353 | case True | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 354 | then have "e > 0" using cs order.order_iff_strict xa_cball xa_y by fastforce | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 355 | then show "y \<in> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 356 | using \<open>x = y \<close> by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 357 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 358 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 359 | have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 360 | unfolding dist_norm | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 361 | using \<open>d>0\<close> norm_ge_zero[of "y - x"] \<open>x \<noteq> y\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 362 | then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 363 | using d as(1)[unfolded subset_eq] by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 364 | have "y - x \<noteq> 0" using \<open>x \<noteq> y\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 365 | hence **:"d / (2 * norm (y - x)) > 0" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 366 | unfolding zero_less_norm_iff[symmetric] using \<open>d>0\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 367 | have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 368 | norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 369 | by (auto simp: dist_norm algebra_simps) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 370 | also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 371 | by (auto simp: algebra_simps) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 372 | also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 373 | using ** by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 374 | also have "\<dots> = (dist y x) + d/2" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 375 | using ** by (auto simp: distrib_right dist_norm) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 376 | finally have "e \<ge> dist x y +d/2" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 377 | using *[unfolded mem_cball] by (auto simp: dist_commute) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 378 | then show "y \<in> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 379 | unfolding mem_ball using \<open>d>0\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 380 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 381 | } | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 382 | then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 383 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 384 | ultimately show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 385 | using interior_unique[of "ball x e" "cball x e"] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 386 | using open_ball[of x e] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 387 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 388 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 389 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 390 | lemma frontier_ball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 391 | fixes a :: "'a::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 392 | shows "0 < e \<Longrightarrow> frontier (ball a e) = sphere a e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 393 | by (force simp: frontier_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 394 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 395 | lemma frontier_cball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 396 |   fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 397 | shows "frontier (cball a e) = sphere a e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 398 | by (force simp: frontier_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 399 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 400 | corollary compact_sphere [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 401 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 402 | shows "compact (sphere a r)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 403 | using compact_frontier [of "cball a r"] by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 404 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 405 | corollary bounded_sphere [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 406 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 407 | shows "bounded (sphere a r)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 408 | by (simp add: compact_imp_bounded) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 409 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 410 | corollary closed_sphere [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 411 |   fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 412 | shows "closed (sphere a r)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 413 | by (simp add: compact_imp_closed) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 414 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 415 | lemma image_add_ball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 416 | fixes a :: "'a::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 417 | shows "(+) b ` ball a r = ball (a+b) r" | 
| 72643 | 418 | proof - | 
| 419 |   { fix x :: 'a
 | |
| 420 | assume "dist (a + b) x < r" | |
| 421 | moreover | |
| 422 | have "b + (x - b) = x" | |
| 423 | by simp | |
| 424 | ultimately have "x \<in> (+) b ` ball a r" | |
| 425 | by (metis add.commute dist_add_cancel image_eqI mem_ball) } | |
| 426 | then show ?thesis | |
| 427 | by (auto simp: add.commute) | |
| 428 | qed | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 429 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 430 | lemma image_add_cball [simp]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 431 | fixes a :: "'a::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 432 | shows "(+) b ` cball a r = cball (a+b) r" | 
| 72643 | 433 | proof - | 
| 434 | have "\<And>x. dist (a + b) x \<le> r \<Longrightarrow> \<exists>y\<in>cball a r. x = b + y" | |
| 435 | by (metis (no_types) add.commute diff_add_cancel dist_add_cancel2 mem_cball) | |
| 436 | then show ?thesis | |
| 437 | by (force simp: add.commute) | |
| 438 | qed | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 439 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 440 | |
| 70136 | 441 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas on Normed Algebras\<close> | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 442 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 443 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 444 | lemma closed_of_nat_image: "closed (of_nat ` A :: 'a::real_normed_algebra_1 set)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 445 | by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_nat) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 446 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 447 | lemma closed_of_int_image: "closed (of_int ` A :: 'a::real_normed_algebra_1 set)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 448 | by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_int) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 449 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 450 | lemma closed_Nats [simp]: "closed (\<nat> :: 'a :: real_normed_algebra_1 set)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 451 | unfolding Nats_def by (rule closed_of_nat_image) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 452 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 453 | lemma closed_Ints [simp]: "closed (\<int> :: 'a :: real_normed_algebra_1 set)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 454 | unfolding Ints_def by (rule closed_of_int_image) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 455 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 456 | lemma closed_subset_Ints: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 457 | fixes A :: "'a :: real_normed_algebra_1 set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 458 | assumes "A \<subseteq> \<int>" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 459 | shows "closed A" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 460 | proof (intro discrete_imp_closed[OF zero_less_one] ballI impI, goal_cases) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 461 | case (1 x y) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 462 | with assms have "x \<in> \<int>" and "y \<in> \<int>" by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 463 | with \<open>dist y x < 1\<close> show "y = x" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 464 | by (auto elim!: Ints_cases simp: dist_of_int) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 465 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 466 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 467 | subsection \<open>Filters\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 468 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 469 | definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" (infixr "indirection" 70) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 470 |   where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 471 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 472 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 473 | subsection \<open>Trivial Limits\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 474 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 475 | lemma trivial_limit_at_infinity: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 476 |   "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)"
 | 
| 72643 | 477 | proof - | 
| 478 | obtain x::'a where "x \<noteq> 0" | |
| 479 | by (meson perfect_choose_dist zero_less_one) | |
| 480 | then have "b \<le> norm ((b / norm x) *\<^sub>R x)" for b | |
| 481 | by simp | |
| 482 | then show ?thesis | |
| 483 | unfolding trivial_limit_def eventually_at_infinity | |
| 484 | by blast | |
| 485 | qed | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 486 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 487 | lemma at_within_ball_bot_iff: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 488 |   fixes x y :: "'a::{real_normed_vector,perfect_space}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 489 | shows "at x within ball y r = bot \<longleftrightarrow> (r=0 \<or> x \<notin> cball y r)" | 
| 72643 | 490 | unfolding trivial_limit_within | 
| 491 | by (metis (no_types) cball_empty equals0D islimpt_ball less_linear) | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 492 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 493 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 494 | subsection \<open>Limits\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 495 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 496 | proposition Lim_at_infinity: "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 497 | by (auto simp: tendsto_iff eventually_at_infinity) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 498 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 499 | corollary Lim_at_infinityI [intro?]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 500 | assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>B. \<forall>x. norm x \<ge> B \<longrightarrow> dist (f x) l \<le> e" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 501 | shows "(f \<longlongrightarrow> l) at_infinity" | 
| 72643 | 502 | proof - | 
| 503 | have "\<And>e. e > 0 \<Longrightarrow> \<exists>B. \<forall>x. norm x \<ge> B \<longrightarrow> dist (f x) l < e" | |
| 504 | by (meson assms dense le_less_trans) | |
| 505 | then show ?thesis | |
| 506 | using Lim_at_infinity by blast | |
| 507 | qed | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 508 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 509 | lemma Lim_transform_within_set_eq: | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 510 | fixes a :: "'a::metric_space" and l :: "'b::metric_space" | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 511 | shows "eventually (\<lambda>x. x \<in> S \<longleftrightarrow> x \<in> T) (at a) | 
| 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 512 | \<Longrightarrow> ((f \<longlongrightarrow> l) (at a within S) \<longleftrightarrow> (f \<longlongrightarrow> l) (at a within T))" | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 513 | by (force intro: Lim_transform_within_set elim: eventually_mono) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 514 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 515 | lemma Lim_null: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 516 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 517 | shows "(f \<longlongrightarrow> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 518 | by (simp add: Lim dist_norm) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 519 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 520 | lemma Lim_null_comparison: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 521 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 522 | assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 523 | shows "(f \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 524 | using assms(2) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 525 | proof (rule metric_tendsto_imp_tendsto) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 526 | show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 527 | using assms(1) by (rule eventually_mono) (simp add: dist_norm) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 528 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 529 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 530 | lemma Lim_transform_bound: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 531 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 532 | and g :: "'a \<Rightarrow> 'c::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 533 | assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 534 | and "(g \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 535 | shows "(f \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 536 | using assms(1) tendsto_norm_zero [OF assms(2)] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 537 | by (rule Lim_null_comparison) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 538 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 539 | lemma lim_null_mult_right_bounded: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 540 | fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 541 | assumes f: "(f \<longlongrightarrow> 0) F" and g: "eventually (\<lambda>x. norm(g x) \<le> B) F" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 542 | shows "((\<lambda>z. f z * g z) \<longlongrightarrow> 0) F" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 543 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 544 | have "((\<lambda>x. norm (f x) * norm (g x)) \<longlongrightarrow> 0) F" | 
| 72643 | 545 | proof (rule Lim_null_comparison) | 
| 546 | show "\<forall>\<^sub>F x in F. norm (norm (f x) * norm (g x)) \<le> norm (f x) * B" | |
| 547 | by (simp add: eventually_mono [OF g] mult_left_mono) | |
| 548 | show "((\<lambda>x. norm (f x) * B) \<longlongrightarrow> 0) F" | |
| 549 | by (simp add: f tendsto_mult_left_zero tendsto_norm_zero) | |
| 550 | qed | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 551 | then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 552 | by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 553 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 554 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 555 | lemma lim_null_mult_left_bounded: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 556 | fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 557 | assumes g: "eventually (\<lambda>x. norm(g x) \<le> B) F" and f: "(f \<longlongrightarrow> 0) F" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 558 | shows "((\<lambda>z. g z * f z) \<longlongrightarrow> 0) F" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 559 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 560 | have "((\<lambda>x. norm (g x) * norm (f x)) \<longlongrightarrow> 0) F" | 
| 72643 | 561 | proof (rule Lim_null_comparison) | 
| 562 | show "\<forall>\<^sub>F x in F. norm (norm (g x) * norm (f x)) \<le> B * norm (f x)" | |
| 563 | by (simp add: eventually_mono [OF g] mult_right_mono) | |
| 564 | show "((\<lambda>x. B * norm (f x)) \<longlongrightarrow> 0) F" | |
| 565 | by (simp add: f tendsto_mult_right_zero tendsto_norm_zero) | |
| 566 | qed | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 567 | then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 568 | by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 569 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 570 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 571 | lemma lim_null_scaleR_bounded: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 572 | assumes f: "(f \<longlongrightarrow> 0) net" and gB: "eventually (\<lambda>a. f a = 0 \<or> norm(g a) \<le> B) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 573 | shows "((\<lambda>n. f n *\<^sub>R g n) \<longlongrightarrow> 0) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 574 | proof | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 575 | fix \<epsilon>::real | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 576 | assume "0 < \<epsilon>" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 577 | then have B: "0 < \<epsilon> / (abs B + 1)" by simp | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 578 | have *: "\<bar>f x\<bar> * norm (g x) < \<epsilon>" if f: "\<bar>f x\<bar> * (\<bar>B\<bar> + 1) < \<epsilon>" and g: "norm (g x) \<le> B" for x | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 579 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 580 | have "\<bar>f x\<bar> * norm (g x) \<le> \<bar>f x\<bar> * B" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 581 | by (simp add: mult_left_mono g) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 582 | also have "\<dots> \<le> \<bar>f x\<bar> * (\<bar>B\<bar> + 1)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 583 | by (simp add: mult_left_mono) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 584 | also have "\<dots> < \<epsilon>" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 585 | by (rule f) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 586 | finally show ?thesis . | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 587 | qed | 
| 72643 | 588 | have "\<And>x. \<lbrakk>\<bar>f x\<bar> < \<epsilon> / (\<bar>B\<bar> + 1); norm (g x) \<le> B\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> * norm (g x) < \<epsilon>" | 
| 589 | by (simp add: "*" pos_less_divide_eq) | |
| 590 | then show "\<forall>\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \<epsilon>" | |
| 591 | using \<open>0 < \<epsilon>\<close> by (auto intro: eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB]]) | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 592 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 593 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 594 | lemma Lim_norm_ubound: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 595 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 596 | assumes "\<not>(trivial_limit net)" "(f \<longlongrightarrow> l) net" "eventually (\<lambda>x. norm(f x) \<le> e) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 597 | shows "norm(l) \<le> e" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 598 | using assms by (fast intro: tendsto_le tendsto_intros) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 599 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 600 | lemma Lim_norm_lbound: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 601 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 602 | assumes "\<not> trivial_limit net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 603 | and "(f \<longlongrightarrow> l) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 604 | and "eventually (\<lambda>x. e \<le> norm (f x)) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 605 | shows "e \<le> norm l" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 606 | using assms by (fast intro: tendsto_le tendsto_intros) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 607 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 608 | text\<open>Limit under bilinear function\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 609 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 610 | lemma Lim_bilinear: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 611 | assumes "(f \<longlongrightarrow> l) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 612 | and "(g \<longlongrightarrow> m) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 613 | and "bounded_bilinear h" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 614 | shows "((\<lambda>x. h (f x) (g x)) \<longlongrightarrow> (h l m)) net" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 615 | using \<open>bounded_bilinear h\<close> \<open>(f \<longlongrightarrow> l) net\<close> \<open>(g \<longlongrightarrow> m) net\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 616 | by (rule bounded_bilinear.tendsto) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 617 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 618 | lemma Lim_at_zero: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 619 | fixes a :: "'a::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 620 | and l :: "'b::topological_space" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 621 | shows "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) \<longlongrightarrow> l) (at 0)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 622 | using LIM_offset_zero LIM_offset_zero_cancel .. | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 623 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 624 | |
| 70136 | 625 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Limit Point of Filter\<close> | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 626 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 627 | lemma netlimit_at_vector: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 628 | fixes a :: "'a::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 629 | shows "netlimit (at a) = a" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 630 | proof (cases "\<exists>x. x \<noteq> a") | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 631 | case True then obtain x where x: "x \<noteq> a" .. | 
| 72643 | 632 | have "\<And>d. 0 < d \<Longrightarrow> \<exists>x. x \<noteq> a \<and> norm (x - a) < d" | 
| 633 | by (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) (simp add: norm_sgn sgn_zero_iff x) | |
| 634 | then have "\<not> trivial_limit (at a)" | |
| 635 | by (auto simp: trivial_limit_def eventually_at dist_norm) | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 636 | then show ?thesis | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 637 | by (rule Lim_ident_at [of a UNIV]) | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 638 | qed simp | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 639 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 640 | subsection \<open>Boundedness\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 641 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 642 | lemma continuous_on_closure_norm_le: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 643 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 644 | assumes "continuous_on (closure s) f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 645 | and "\<forall>y \<in> s. norm(f y) \<le> b" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 646 | and "x \<in> (closure s)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 647 | shows "norm (f x) \<le> b" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 648 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 649 | have *: "f ` s \<subseteq> cball 0 b" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 650 | using assms(2)[unfolded mem_cball_0[symmetric]] by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 651 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 652 | by (meson "*" assms(1) assms(3) closed_cball image_closure_subset image_subset_iff mem_cball_0) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 653 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 654 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 655 | lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)" | 
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70346diff
changeset | 656 | unfolding bounded_iff | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70346diff
changeset | 657 | by (meson less_imp_le not_le order_trans zero_less_one) | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 658 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 659 | lemma bounded_pos_less: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x < b)" | 
| 72643 | 660 | by (metis bounded_pos le_less_trans less_imp_le linordered_field_no_ub) | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 661 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 662 | lemma Bseq_eq_bounded: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 663 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 664 | shows "Bseq f \<longleftrightarrow> bounded (range f)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 665 | unfolding Bseq_def bounded_pos by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 666 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 667 | lemma bounded_linear_image: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 668 | assumes "bounded S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 669 | and "bounded_linear f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 670 | shows "bounded (f ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 671 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 672 | from assms(1) obtain b where "b > 0" and b: "\<forall>x\<in>S. norm x \<le> b" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 673 | unfolding bounded_pos by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 674 | from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 675 | using bounded_linear.pos_bounded by (auto simp: ac_simps) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 676 | show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 677 | unfolding bounded_pos | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 678 | proof (intro exI, safe) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 679 | show "norm (f x) \<le> B * b" if "x \<in> S" for x | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 680 | by (meson B b less_imp_le mult_left_mono order_trans that) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 681 | qed (use \<open>b > 0\<close> \<open>B > 0\<close> in auto) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 682 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 683 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 684 | lemma bounded_scaling: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 685 | fixes S :: "'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 686 | shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)" | 
| 72643 | 687 | by (simp add: bounded_linear_image bounded_linear_scaleR_right) | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 688 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 689 | lemma bounded_scaleR_comp: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 690 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 691 | assumes "bounded (f ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 692 | shows "bounded ((\<lambda>x. r *\<^sub>R f x) ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 693 | using bounded_scaling[of "f ` S" r] assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 694 | by (auto simp: image_image) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 695 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 696 | lemma bounded_translation: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 697 | fixes S :: "'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 698 | assumes "bounded S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 699 | shows "bounded ((\<lambda>x. a + x) ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 700 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 701 | from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 702 | unfolding bounded_pos by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 703 |   {
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 704 | fix x | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 705 | assume "x \<in> S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 706 | then have "norm (a + x) \<le> b + norm a" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 707 | using norm_triangle_ineq[of a x] b by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 708 | } | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 709 | then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 710 | unfolding bounded_pos | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 711 | using norm_ge_zero[of a] b(1) and add_strict_increasing[of b 0 "norm a"] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 712 | by (auto intro!: exI[of _ "b + norm a"]) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 713 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 714 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 715 | lemma bounded_translation_minus: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 716 | fixes S :: "'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 717 | shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. x - a) ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 718 | using bounded_translation [of S "-a"] by simp | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 719 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 720 | lemma bounded_uminus [simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 721 | fixes X :: "'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 722 | shows "bounded (uminus ` X) \<longleftrightarrow> bounded X" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 723 | by (auto simp: bounded_def dist_norm; rule_tac x="-x" in exI; force simp: add.commute norm_minus_commute) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 724 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 725 | lemma uminus_bounded_comp [simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 726 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 727 | shows "bounded ((\<lambda>x. - f x) ` S) \<longleftrightarrow> bounded (f ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 728 | using bounded_uminus[of "f ` S"] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 729 | by (auto simp: image_image) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 730 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 731 | lemma bounded_plus_comp: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 732 | fixes f g::"'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 733 | assumes "bounded (f ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 734 | assumes "bounded (g ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 735 | shows "bounded ((\<lambda>x. f x + g x) ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 736 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 737 |   {
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 738 | fix B C | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 739 | assume "\<And>x. x\<in>S \<Longrightarrow> norm (f x) \<le> B" "\<And>x. x\<in>S \<Longrightarrow> norm (g x) \<le> C" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 740 | then have "\<And>x. x \<in> S \<Longrightarrow> norm (f x + g x) \<le> B + C" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 741 | by (auto intro!: norm_triangle_le add_mono) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 742 | } then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 743 | using assms by (fastforce simp: bounded_iff) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 744 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 745 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 746 | lemma bounded_plus: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 747 | fixes S ::"'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 748 | assumes "bounded S" "bounded T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 749 | shows "bounded ((\<lambda>(x,y). x + y) ` (S \<times> T))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 750 | using bounded_plus_comp [of fst "S \<times> T" snd] assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 751 | by (auto simp: split_def split: if_split_asm) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 752 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 753 | lemma bounded_minus_comp: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 754 | "bounded (f ` S) \<Longrightarrow> bounded (g ` S) \<Longrightarrow> bounded ((\<lambda>x. f x - g x) ` S)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 755 | for f g::"'a \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 756 | using bounded_plus_comp[of "f" S "\<lambda>x. - g x"] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 757 | by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 758 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 759 | lemma bounded_minus: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 760 | fixes S ::"'a::real_normed_vector set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 761 | assumes "bounded S" "bounded T" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 762 | shows "bounded ((\<lambda>(x,y). x - y) ` (S \<times> T))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 763 | using bounded_minus_comp [of fst "S \<times> T" snd] assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 764 | by (auto simp: split_def split: if_split_asm) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 765 | |
| 73791 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 766 | lemma bounded_sums: | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 767 | fixes S :: "'a::real_normed_vector set" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 768 | assumes "bounded S" and "bounded T" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 769 |   shows "bounded (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 770 | using assms by (simp add: bounded_iff) (meson norm_triangle_mono) | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 771 | |
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 772 | lemma bounded_differences: | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 773 | fixes S :: "'a::real_normed_vector set" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 774 | assumes "bounded S" and "bounded T" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 775 |   shows "bounded (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 776 | using assms by (simp add: bounded_iff) (meson add_mono norm_triangle_le_diff) | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 777 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 778 | lemma not_bounded_UNIV[simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 779 |   "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 780 | proof (auto simp: bounded_pos not_le) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 781 | obtain x :: 'a where "x \<noteq> 0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 782 | using perfect_choose_dist [OF zero_less_one] by fast | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 783 | fix b :: real | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 784 | assume b: "b >0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 785 | have b1: "b +1 \<ge> 0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 786 | using b by simp | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 787 | with \<open>x \<noteq> 0\<close> have "b < norm (scaleR (b + 1) (sgn x))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 788 | by (simp add: norm_sgn) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 789 | then show "\<exists>x::'a. b < norm x" .. | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 790 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 791 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 792 | corollary cobounded_imp_unbounded: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 793 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 794 | shows "bounded (- S) \<Longrightarrow> \<not> bounded S" | 
| 71633 | 795 | using bounded_Un [of S "-S"] by (simp) | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 796 | |
| 70136 | 797 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations among convergence and absolute convergence for power series\<close> | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 798 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 799 | lemma summable_imp_bounded: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 800 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 801 | shows "summable f \<Longrightarrow> bounded (range f)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 802 | by (frule summable_LIMSEQ_zero) (simp add: convergent_imp_bounded) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 803 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 804 | lemma summable_imp_sums_bounded: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 805 |    "summable f \<Longrightarrow> bounded (range (\<lambda>n. sum f {..<n}))"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 806 | by (auto simp: summable_def sums_def dest: convergent_imp_bounded) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 807 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 808 | lemma power_series_conv_imp_absconv_weak: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 809 |   fixes a:: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}" and w :: 'a
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 810 | assumes sum: "summable (\<lambda>n. a n * z ^ n)" and no: "norm w < norm z" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 811 | shows "summable (\<lambda>n. of_real(norm(a n)) * w ^ n)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 812 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 813 | obtain M where M: "\<And>x. norm (a x * z ^ x) \<le> M" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 814 | using summable_imp_bounded [OF sum] by (force simp: bounded_iff) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 815 | show ?thesis | 
| 72643 | 816 | proof (rule series_comparison_complex) | 
| 817 | have "\<And>n. norm (a n) * norm z ^ n \<le> M" | |
| 818 | by (metis (no_types) M norm_mult norm_power) | |
| 819 | then show "summable (\<lambda>n. complex_of_real (norm (a n) * norm w ^ n))" | |
| 820 | using Abel_lemma no norm_ge_zero summable_of_real by blast | |
| 821 | qed (auto simp: norm_mult norm_power) | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 822 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 823 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 824 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 825 | subsection \<open>Normed spaces with the Heine-Borel property\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 826 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 827 | lemma not_compact_UNIV[simp]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 828 |   fixes s :: "'a::{real_normed_vector,perfect_space,heine_borel} set"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 829 | shows "\<not> compact (UNIV::'a set)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 830 | by (simp add: compact_eq_bounded_closed) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 831 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 832 | lemma not_compact_space_euclideanreal [simp]: "\<not> compact_space euclideanreal" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 833 | by (simp add: compact_space_def) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 834 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 835 | text\<open>Representing sets as the union of a chain of compact sets.\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 836 | lemma closed_Union_compact_subsets: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 837 |   fixes S :: "'a::{heine_borel,real_normed_vector} set"
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 838 | assumes "closed S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 839 | obtains F where "\<And>n. compact(F n)" "\<And>n. F n \<subseteq> S" "\<And>n. F n \<subseteq> F(Suc n)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 840 | "(\<Union>n. F n) = S" "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. K \<subseteq> F n" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 841 | proof | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 842 | show "compact (S \<inter> cball 0 (of_nat n))" for n | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 843 | using assms compact_eq_bounded_closed by auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 844 | next | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 845 | show "(\<Union>n. S \<inter> cball 0 (real n)) = S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 846 | by (auto simp: real_arch_simple) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 847 | next | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 848 | fix K :: "'a set" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 849 | assume "compact K" "K \<subseteq> S" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 850 | then obtain N where "K \<subseteq> cball 0 N" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 851 | by (meson bounded_pos mem_cball_0 compact_imp_bounded subsetI) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 852 | then show "\<exists>N. \<forall>n\<ge>N. K \<subseteq> S \<inter> cball 0 (real n)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 853 | by (metis of_nat_le_iff Int_subset_iff \<open>K \<subseteq> S\<close> real_arch_simple subset_cball subset_trans) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 854 | qed auto | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 855 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 856 | subsection \<open>Intersecting chains of compact sets and the Baire property\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 857 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 858 | proposition bounded_closed_chain: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 859 | fixes \<F> :: "'a::heine_borel set set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 860 |   assumes "B \<in> \<F>" "bounded B" and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" and "{} \<notin> \<F>"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 861 | and chain: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 862 |     shows "\<Inter>\<F> \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 863 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 864 |   have "B \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 865 | proof (rule compact_imp_fip) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 866 | show "compact B" "\<And>T. T \<in> \<F> \<Longrightarrow> closed T" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 867 | by (simp_all add: assms compact_eq_bounded_closed) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 868 |     show "\<lbrakk>finite \<G>; \<G> \<subseteq> \<F>\<rbrakk> \<Longrightarrow> B \<inter> \<Inter>\<G> \<noteq> {}" for \<G>
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 869 | proof (induction \<G> rule: finite_induct) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 870 | case empty | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 871 | with assms show ?case by force | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 872 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 873 | case (insert U \<G>) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 874 |       then have "U \<in> \<F>" and ne: "B \<inter> \<Inter>\<G> \<noteq> {}" by auto
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 875 | then consider "B \<subseteq> U" | "U \<subseteq> B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 876 | using \<open>B \<in> \<F>\<close> chain by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 877 | then show ?case | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 878 | proof cases | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 879 | case 1 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 880 | then show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 881 | using Int_left_commute ne by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 882 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 883 | case 2 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 884 |           have "U \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 885 |             using \<open>U \<in> \<F>\<close> \<open>{} \<notin> \<F>\<close> by blast
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 886 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 887 | have False if "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. x \<notin> Y" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 888 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 889 | have "\<And>x. x \<in> U \<Longrightarrow> \<exists>Y\<in>\<G>. Y \<subseteq> U" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 890 | by (metis chain contra_subsetD insert.prems insert_subset that) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 891 | then obtain Y where "Y \<in> \<G>" "Y \<subseteq> U" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 892 |               by (metis all_not_in_conv \<open>U \<noteq> {}\<close>)
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 893 | moreover obtain x where "x \<in> \<Inter>\<G>" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 894 | by (metis Int_emptyI ne) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 895 | ultimately show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 896 | by (metis Inf_lower subset_eq that) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 897 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 898 | with 2 show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 899 | by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 900 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 901 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 902 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 903 | then show ?thesis by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 904 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 905 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 906 | corollary compact_chain: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 907 | fixes \<F> :: "'a::heine_borel set set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 908 |   assumes "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" "{} \<notin> \<F>"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 909 | "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 910 |     shows "\<Inter> \<F> \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 911 | proof (cases "\<F> = {}")
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 912 | case True | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 913 | then show ?thesis by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 914 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 915 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 916 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 917 | by (metis False all_not_in_conv assms compact_imp_bounded compact_imp_closed bounded_closed_chain) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 918 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 919 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 920 | lemma compact_nest: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 921 | fixes F :: "'a::linorder \<Rightarrow> 'b::heine_borel set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 922 |   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"
 | 
| 69745 | 923 |   shows "\<Inter>(range F) \<noteq> {}"
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 924 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 925 | have *: "\<And>S T. S \<in> range F \<and> T \<in> range F \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 926 | by (metis mono image_iff le_cases) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 927 | show ?thesis | 
| 72643 | 928 | using F by (intro compact_chain [OF _ _ *]; blast dest: *) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 929 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 930 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 931 | text\<open>The Baire property of dense sets\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 932 | theorem Baire: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 933 |   fixes S::"'a::{real_normed_vector,heine_borel} set"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 934 | assumes "closed S" "countable \<G>" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 935 | and ope: "\<And>T. T \<in> \<G> \<Longrightarrow> openin (top_of_set S) T \<and> S \<subseteq> closure T" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 936 | shows "S \<subseteq> closure(\<Inter>\<G>)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 937 | proof (cases "\<G> = {}")
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 938 | case True | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 939 | then show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 940 | using closure_subset by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 941 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 942 | let ?g = "from_nat_into \<G>" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 943 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 944 | then have gin: "?g n \<in> \<G>" for n | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 945 | by (simp add: from_nat_into) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 946 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 947 | proof (clarsimp simp: closure_approachable) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 948 | fix x and e::real | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 949 | assume "x \<in> S" "0 < e" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 950 | obtain TF where opeF: "\<And>n. openin (top_of_set S) (TF n)" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 951 |                and ne: "\<And>n. TF n \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 952 | and subg: "\<And>n. S \<inter> closure(TF n) \<subseteq> ?g n" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 953 | and subball: "\<And>n. closure(TF n) \<subseteq> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 954 | and decr: "\<And>n. TF(Suc n) \<subseteq> TF n" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 955 | proof - | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 956 |       have *: "\<exists>Y. (openin (top_of_set S) Y \<and> Y \<noteq> {} \<and>
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 957 | S \<inter> closure Y \<subseteq> ?g n \<and> closure Y \<subseteq> ball x e) \<and> Y \<subseteq> U" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 958 |         if opeU: "openin (top_of_set S) U" and "U \<noteq> {}" and cloU: "closure U \<subseteq> ball x e" for U n
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 959 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 960 | obtain T where T: "open T" "U = T \<inter> S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 961 | using \<open>openin (top_of_set S) U\<close> by (auto simp: openin_subtopology) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 962 |         with \<open>U \<noteq> {}\<close> have "T \<inter> closure (?g n) \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 963 | using gin ope by fastforce | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 964 |         then have "T \<inter> ?g n \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 965 | using \<open>open T\<close> open_Int_closure_eq_empty by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 966 | then obtain y where "y \<in> U" "y \<in> ?g n" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 967 | using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 968 | moreover have "openin (top_of_set S) (U \<inter> ?g n)" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 969 | using gin ope opeU by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 970 | 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" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 971 | by (force simp: openin_contains_ball) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 972 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 973 | proof (intro exI conjI) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 974 | show "openin (top_of_set S) (S \<inter> ball y (d/2))" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 975 | by (simp add: openin_open_Int) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 976 |           show "S \<inter> ball y (d/2) \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 977 | using \<open>0 < d\<close> \<open>y \<in> U\<close> opeU openin_imp_subset by fastforce | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 978 | have "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> closure (ball y (d/2))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 979 | using closure_mono by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 980 | also have "... \<subseteq> ?g n" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 981 | using \<open>d > 0\<close> d by force | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 982 | finally show "S \<inter> closure (S \<inter> ball y (d/2)) \<subseteq> ?g n" . | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 983 | have "closure (S \<inter> ball y (d/2)) \<subseteq> S \<inter> ball y d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 984 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 985 | have "closure (ball y (d/2)) \<subseteq> ball y d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 986 | using \<open>d > 0\<close> by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 987 | then have "closure (S \<inter> ball y (d/2)) \<subseteq> ball y d" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 988 | by (meson closure_mono inf.cobounded2 subset_trans) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 989 | then show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 990 | by (simp add: \<open>closed S\<close> closure_minimal) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 991 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 992 | also have "... \<subseteq> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 993 | using cloU closure_subset d by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 994 | finally show "closure (S \<inter> ball y (d/2)) \<subseteq> ball x e" . | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 995 | show "S \<inter> ball y (d/2) \<subseteq> U" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 996 | using ball_divide_subset_numeral d by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 997 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 998 | qed | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 999 |       let ?\<Phi> = "\<lambda>n X. openin (top_of_set S) X \<and> X \<noteq> {} \<and>
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1000 | S \<inter> closure X \<subseteq> ?g n \<and> closure X \<subseteq> ball x e" | 
| 72643 | 1001 | have "closure (S \<inter> ball x (e/2)) \<subseteq> closure(ball x (e/2))" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1002 | by (simp add: closure_mono) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1003 | also have "... \<subseteq> ball x e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1004 | using \<open>e > 0\<close> by auto | 
| 72643 | 1005 | finally have "closure (S \<inter> ball x (e/2)) \<subseteq> ball x e" . | 
| 1006 |       moreover have"openin (top_of_set S) (S \<inter> ball x (e/2))" "S \<inter> ball x (e/2) \<noteq> {}"
 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1007 | using \<open>0 < e\<close> \<open>x \<in> S\<close> by auto | 
| 72643 | 1008 | ultimately obtain Y where Y: "?\<Phi> 0 Y \<and> Y \<subseteq> S \<inter> ball x (e/2)" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1009 | using * [of "S \<inter> ball x (e/2)" 0] by metis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1010 | show thesis | 
| 72643 | 1011 | proof (rule exE [OF dependent_nat_choice]) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1012 | show "\<exists>x. ?\<Phi> 0 x" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1013 | using Y by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1014 | show "\<exists>Y. ?\<Phi> (Suc n) Y \<and> Y \<subseteq> X" if "?\<Phi> n X" for X n | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1015 | using that by (blast intro: *) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1016 | qed (use that in metis) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1017 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1018 |     have "(\<Inter>n. S \<inter> closure (TF n)) \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1019 | proof (rule compact_nest) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1020 | show "\<And>n. compact (S \<inter> closure (TF n))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1021 | by (metis closed_closure subball bounded_subset_ballI compact_eq_bounded_closed closed_Int_compact [OF \<open>closed S\<close>]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1022 |       show "\<And>n. S \<inter> closure (TF n) \<noteq> {}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1023 | by (metis Int_absorb1 opeF \<open>closed S\<close> closure_eq_empty closure_minimal ne openin_imp_subset) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1024 | show "\<And>m n. m \<le> n \<Longrightarrow> S \<inter> closure (TF n) \<subseteq> S \<inter> closure (TF m)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1025 | by (meson closure_mono decr dual_order.refl inf_mono lift_Suc_antimono_le) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1026 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1027 |     moreover have "(\<Inter>n. S \<inter> closure (TF n)) \<subseteq> {y \<in> \<Inter>\<G>. dist y x < e}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1028 | proof (clarsimp, intro conjI) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1029 | fix y | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1030 | assume "y \<in> S" and y: "\<forall>n. y \<in> closure (TF n)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1031 | then show "\<forall>T\<in>\<G>. y \<in> T" | 
| 69712 | 1032 | by (metis Int_iff from_nat_into_surj [OF \<open>countable \<G>\<close>] subsetD subg) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1033 | show "dist y x < e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1034 | by (metis y dist_commute mem_ball subball subsetCE) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1035 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1036 | ultimately show "\<exists>y \<in> \<Inter>\<G>. dist y x < e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1037 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1038 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1039 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1040 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1041 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1042 | subsection \<open>Continuity\<close> | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1043 | |
| 70136 | 1044 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for uniform continuity\<close> | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1045 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1046 | lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1047 | fixes g :: "_::metric_space \<Rightarrow> _" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1048 | assumes "uniformly_continuous_on s g" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1049 | shows "uniformly_continuous_on s (\<lambda>x. f (g x))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1050 | using assms unfolding uniformly_continuous_on_sequentially | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1051 | unfolding dist_norm tendsto_norm_zero_iff diff[symmetric] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1052 | by (auto intro: tendsto_zero) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1053 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1054 | lemma uniformly_continuous_on_dist[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1055 | fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1056 | assumes "uniformly_continuous_on s f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1057 | and "uniformly_continuous_on s g" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1058 | shows "uniformly_continuous_on s (\<lambda>x. dist (f x) (g x))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1059 | proof - | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1060 |   {
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1061 | fix a b c d :: 'b | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1062 | have "\<bar>dist a b - dist c d\<bar> \<le> dist a c + dist b d" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1063 | using dist_triangle2 [of a b c] dist_triangle2 [of b c d] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1064 | using dist_triangle3 [of c d a] dist_triangle [of a d b] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1065 | by arith | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1066 | } note le = this | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1067 |   {
 | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1068 | fix x y | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1069 | assume f: "(\<lambda>n. dist (f (x n)) (f (y n))) \<longlonglongrightarrow> 0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1070 | assume g: "(\<lambda>n. dist (g (x n)) (g (y n))) \<longlonglongrightarrow> 0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1071 | have "(\<lambda>n. \<bar>dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\<bar>) \<longlonglongrightarrow> 0" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1072 | by (rule Lim_transform_bound [OF _ tendsto_add_zero [OF f g]], | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1073 | simp add: le) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1074 | } | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1075 | then show ?thesis | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1076 | using assms unfolding uniformly_continuous_on_sequentially | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1077 | unfolding dist_real_def by simp | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1078 | qed | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1079 | |
| 71167 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1080 | lemma uniformly_continuous_on_cmul_right [continuous_intros]: | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1081 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1082 | shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. f x * c)" | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1083 | using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1084 | |
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1085 | lemma uniformly_continuous_on_cmul_left[continuous_intros]: | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1086 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1087 | assumes "uniformly_continuous_on s f" | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1088 | shows "uniformly_continuous_on s (\<lambda>x. c * f x)" | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1089 | by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) | 
| 
b4d409c65a76
Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 1090 | |
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1091 | lemma uniformly_continuous_on_norm[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1092 | fixes f :: "'a :: metric_space \<Rightarrow> 'b :: real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1093 | assumes "uniformly_continuous_on s f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1094 | shows "uniformly_continuous_on s (\<lambda>x. norm (f x))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1095 | unfolding norm_conv_dist using assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1096 | by (intro uniformly_continuous_on_dist uniformly_continuous_on_const) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1097 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1098 | lemma uniformly_continuous_on_cmul[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1099 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1100 | assumes "uniformly_continuous_on s f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1101 | shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1102 | using bounded_linear_scaleR_right assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1103 | by (rule bounded_linear.uniformly_continuous_on) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1104 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1105 | lemma dist_minus: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1106 | fixes x y :: "'a::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1107 | shows "dist (- x) (- y) = dist x y" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1108 | unfolding dist_norm minus_diff_minus norm_minus_cancel .. | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1109 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1110 | lemma uniformly_continuous_on_minus[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1111 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1112 | shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1113 | unfolding uniformly_continuous_on_def dist_minus . | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1114 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1115 | lemma uniformly_continuous_on_add[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1116 | fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1117 | assumes "uniformly_continuous_on s f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1118 | and "uniformly_continuous_on s g" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1119 | shows "uniformly_continuous_on s (\<lambda>x. f x + g x)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1120 | using assms | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1121 | unfolding uniformly_continuous_on_sequentially | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1122 | unfolding dist_norm tendsto_norm_zero_iff add_diff_add | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1123 | by (auto intro: tendsto_add_zero) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1124 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1125 | lemma uniformly_continuous_on_diff[continuous_intros]: | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1126 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1127 | assumes "uniformly_continuous_on s f" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1128 | and "uniformly_continuous_on s g" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1129 | shows "uniformly_continuous_on s (\<lambda>x. f x - g x)" | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1130 | using assms uniformly_continuous_on_add [of s f "- g"] | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1131 | by (simp add: fun_Compl_def uniformly_continuous_on_minus) | 
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1132 | |
| 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: diff
changeset | 1133 | |
| 70136 | 1134 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic Preserves Topological Properties\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1135 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1136 | lemma open_scaling[intro]: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1137 | fixes s :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1138 | assumes "c \<noteq> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1139 | and "open s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1140 | shows "open((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1141 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1142 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1143 | fix x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1144 | assume "x \<in> s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1145 | then obtain e where "e>0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1146 | and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1147 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1148 | have "e * \<bar>c\<bar> > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1149 | using assms(1)[unfolded zero_less_abs_iff[symmetric]] \<open>e>0\<close> by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1150 | moreover | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1151 |     {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1152 | fix y | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1153 | assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>" | 
| 72643 | 1154 | then have "norm (c *\<^sub>R ((1 / c) *\<^sub>R y - x)) < e * norm c" | 
| 1155 | by (simp add: \<open>c \<noteq> 0\<close> dist_norm scale_right_diff_distrib) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1156 | then have "norm ((1 / c) *\<^sub>R y - x) < e" | 
| 72643 | 1157 | by (simp add: \<open>c \<noteq> 0\<close>) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1158 | then have "y \<in> (*\<^sub>R) c ` s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1159 | using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "(*\<^sub>R) c"] | 
| 72643 | 1160 | by (simp add: \<open>c \<noteq> 0\<close> dist_norm e) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1161 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1162 | ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> (*\<^sub>R) c ` s" | 
| 72643 | 1163 | by (rule_tac x="e * \<bar>c\<bar>" in exI, auto) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1164 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1165 | then show ?thesis unfolding open_dist by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1166 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1167 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1168 | lemma minus_image_eq_vimage: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1169 | fixes A :: "'a::ab_group_add set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1170 | shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1171 | by (auto intro!: image_eqI [where f="\<lambda>x. - x"]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1172 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1173 | lemma open_negations: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1174 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1175 | shows "open S \<Longrightarrow> open ((\<lambda>x. - x) ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1176 | using open_scaling [of "- 1" S] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1177 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1178 | lemma open_translation: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1179 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1180 | assumes "open S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1181 | shows "open((\<lambda>x. a + x) ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1182 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1183 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1184 | fix x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1185 | have "continuous (at x) (\<lambda>x. x - a)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1186 | by (intro continuous_diff continuous_ident continuous_const) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1187 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1188 |   moreover have "{x. x - a \<in> S} = (+) a ` S"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1189 | by force | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1190 | ultimately show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1191 | by (metis assms continuous_open_vimage vimage_def) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1192 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1193 | |
| 70999 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1194 | lemma open_translation_subtract: | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1195 | fixes S :: "'a::real_normed_vector set" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1196 | assumes "open S" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1197 | shows "open ((\<lambda>x. x - a) ` S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1198 | using assms open_translation [of S "- a"] by (simp cong: image_cong_simp) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1199 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1200 | lemma open_neg_translation: | 
| 70999 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1201 | fixes S :: "'a::real_normed_vector set" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1202 | assumes "open S" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70817diff
changeset | 1203 | shows "open((\<lambda>x. a - x) ` S)" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1204 | using open_translation[OF open_negations[OF assms], of a] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1205 | by (auto simp: image_image) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1206 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1207 | lemma open_affinity: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1208 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1209 | assumes "open S" "c \<noteq> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1210 | shows "open ((\<lambda>x. a + c *\<^sub>R x) ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1211 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1212 | have *: "(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1213 | unfolding o_def .. | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1214 | have "(+) a ` (*\<^sub>R) c ` S = ((+) a \<circ> (*\<^sub>R) c) ` S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1215 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1216 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1217 | using assms open_translation[of "(*\<^sub>R) c ` S" a] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1218 | unfolding * | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1219 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1220 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1221 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1222 | lemma interior_translation: | 
| 69661 | 1223 | "interior ((+) a ` S) = (+) a ` (interior S)" for S :: "'a::real_normed_vector set" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1224 | proof (rule set_eqI, rule) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1225 | fix x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1226 | assume "x \<in> interior ((+) a ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1227 | then obtain e where "e > 0" and e: "ball x e \<subseteq> (+) a ` S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1228 | unfolding mem_interior by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1229 | then have "ball (x - a) e \<subseteq> S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1230 | unfolding subset_eq Ball_def mem_ball dist_norm | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1231 | by (auto simp: diff_diff_eq) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1232 | then show "x \<in> (+) a ` interior S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1233 | unfolding image_iff | 
| 72643 | 1234 | by (metis \<open>0 < e\<close> add.commute diff_add_cancel mem_interior) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1235 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1236 | fix x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1237 | assume "x \<in> (+) a ` interior S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1238 | then obtain y e where "e > 0" and e: "ball y e \<subseteq> S" and y: "x = a + y" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1239 | unfolding image_iff Bex_def mem_interior by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1240 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1241 | fix z | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1242 | have *: "a + y - z = y + a - z" by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1243 | assume "z \<in> ball x e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1244 | then have "z - a \<in> S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1245 | using e[unfolded subset_eq, THEN bspec[where x="z - a"]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1246 | unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1247 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1248 | then have "z \<in> (+) a ` S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1249 | unfolding image_iff by (auto intro!: bexI[where x="z - a"]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1250 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1251 | then have "ball x e \<subseteq> (+) a ` S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1252 | unfolding subset_eq by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1253 | then show "x \<in> interior ((+) a ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1254 | unfolding mem_interior using \<open>e > 0\<close> by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1255 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1256 | |
| 69661 | 1257 | lemma interior_translation_subtract: | 
| 1258 | "interior ((\<lambda>x. x - a) ` S) = (\<lambda>x. x - a) ` interior S" for S :: "'a::real_normed_vector set" | |
| 1259 | using interior_translation [of "- a"] by (simp cong: image_cong_simp) | |
| 1260 | ||
| 1261 | ||
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1262 | lemma compact_scaling: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1263 | fixes s :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1264 | assumes "compact s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1265 | shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1266 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1267 | let ?f = "\<lambda>x. scaleR c x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1268 | have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1269 | show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1270 | using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1271 | using linear_continuous_at[OF *] assms | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1272 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1273 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1274 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1275 | lemma compact_negations: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1276 | fixes s :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1277 | assumes "compact s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1278 | shows "compact ((\<lambda>x. - x) ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1279 | using compact_scaling [OF assms, of "- 1"] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1280 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1281 | lemma compact_sums: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1282 | fixes s t :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1283 | assumes "compact s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1284 | and "compact t" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1285 |   shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1286 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1287 |   have *: "{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
 | 
| 72643 | 1288 | by (fastforce simp: image_iff) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1289 | have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1290 | unfolding continuous_on by (rule ballI) (intro tendsto_intros) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1291 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1292 | unfolding * using compact_continuous_image compact_Times [OF assms] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1293 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1294 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1295 | lemma compact_differences: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1296 | fixes s t :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1297 | assumes "compact s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1298 | and "compact t" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1299 |   shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1300 | proof- | 
| 72643 | 1301 |   have "{x - y | x y. x\<in>s \<and> y \<in> t} = {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
 | 
| 1302 | using diff_conv_add_uminus by force | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1303 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1304 | using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1305 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1306 | |
| 73791 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1307 | lemma compact_sums': | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1308 | fixes S :: "'a::real_normed_vector set" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1309 | assumes "compact S" and "compact T" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1310 |   shows "compact (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1311 | proof - | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1312 |   have "(\<Union>x\<in>S. \<Union>y\<in>T. {x + y}) = {x + y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1313 | by blast | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1314 | then show ?thesis | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1315 | using compact_sums [OF assms] by simp | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1316 | qed | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1317 | |
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1318 | lemma compact_differences': | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1319 | fixes S :: "'a::real_normed_vector set" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1320 | assumes "compact S" and "compact T" | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1321 |   shows "compact (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1322 | proof - | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1323 |   have "(\<Union>x\<in>S. \<Union>y\<in>T. {x - y}) = {x - y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1324 | by blast | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1325 | then show ?thesis | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1326 | using compact_differences [OF assms] by simp | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1327 | qed | 
| 
e10d530f157a
some new and/or varient results about images
 paulson <lp15@cam.ac.uk> parents: 
72643diff
changeset | 1328 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1329 | lemma compact_translation: | 
| 69661 | 1330 | "compact ((+) a ` s)" if "compact s" for s :: "'a::real_normed_vector set" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1331 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1332 |   have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1333 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1334 | then show ?thesis | 
| 69661 | 1335 | using compact_sums [OF that compact_sing [of a]] by auto | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1336 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1337 | |
| 69661 | 1338 | lemma compact_translation_subtract: | 
| 1339 | "compact ((\<lambda>x. x - a) ` s)" if "compact s" for s :: "'a::real_normed_vector set" | |
| 1340 | using that compact_translation [of s "- a"] by (simp cong: image_cong_simp) | |
| 1341 | ||
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1342 | lemma compact_affinity: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1343 | fixes s :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1344 | assumes "compact s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1345 | shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1346 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1347 | have "(+) a ` (*\<^sub>R) c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1348 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1349 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1350 | using compact_translation[OF compact_scaling[OF assms], of a c] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1351 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1352 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1353 | lemma closed_scaling: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1354 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1355 | assumes "closed S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1356 | shows "closed ((\<lambda>x. c *\<^sub>R x) ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1357 | proof (cases "c = 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1358 | case True then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1359 | by (auto simp: image_constant_conv) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1360 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1361 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1362 | from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1363 | by (simp add: continuous_closed_vimage) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1364 | also have "(\<lambda>x. inverse c *\<^sub>R x) -` S = (\<lambda>x. c *\<^sub>R x) ` S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1365 | using \<open>c \<noteq> 0\<close> by (auto elim: image_eqI [rotated]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1366 | finally show ?thesis . | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1367 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1368 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1369 | lemma closed_negations: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1370 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1371 | assumes "closed S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1372 | shows "closed ((\<lambda>x. -x) ` S)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1373 | using closed_scaling[OF assms, of "- 1"] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1374 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1375 | lemma compact_closed_sums: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1376 | fixes S :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1377 | assumes "compact S" and "closed T" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1378 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1379 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1380 |   let ?S = "{x + y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1381 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1382 | fix x l | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1383 | assume as: "\<forall>n. x n \<in> ?S" "(x \<longlongrightarrow> l) sequentially" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1384 | 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" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1385 | using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> S \<and> snd y \<in> T"] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1386 | obtain l' r where "l'\<in>S" and r: "strict_mono r" and lr: "(((\<lambda>n. fst (f n)) \<circ> r) \<longlongrightarrow> l') sequentially" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1387 | using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1388 | have "((\<lambda>n. snd (f (r n))) \<longlongrightarrow> l - l') sequentially" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1389 | using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1390 | unfolding o_def | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1391 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1392 | then have "l - l' \<in> T" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1393 | using assms(2)[unfolded closed_sequential_limits, | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1394 | THEN spec[where x="\<lambda> n. snd (f (r n))"], | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1395 | THEN spec[where x="l - l'"]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1396 | using f(3) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1397 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1398 | then have "l \<in> ?S" | 
| 72643 | 1399 | using \<open>l' \<in> S\<close> by force | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1400 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1401 |   moreover have "?S = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1402 | by force | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1403 | ultimately show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1404 | unfolding closed_sequential_limits | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1405 | by (metis (no_types, lifting)) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1406 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1407 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1408 | lemma closed_compact_sums: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1409 | fixes S T :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1410 | assumes "closed S" "compact T" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1411 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1412 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1413 |   have "(\<Union>x\<in> T. \<Union>y \<in> S. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1414 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1415 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1416 | using compact_closed_sums[OF assms(2,1)] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1417 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1418 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1419 | lemma compact_closed_differences: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1420 | fixes S T :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1421 | assumes "compact S" "closed T" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1422 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1423 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1424 |   have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1425 | by force | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1426 | then show ?thesis | 
| 72643 | 1427 | by (metis assms closed_negations compact_closed_sums) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1428 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1429 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1430 | lemma closed_compact_differences: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1431 | fixes S T :: "'a::real_normed_vector set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1432 | assumes "closed S" "compact T" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1433 |   shows "closed (\<Union>x\<in> S. \<Union>y \<in> T. {x - y})"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1434 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1435 |   have "(\<Union>x\<in> S. \<Union>y \<in> uminus ` T. {x + y}) = {x - y |x y. x \<in> S \<and> y \<in> T}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1436 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1437 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1438 | using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1439 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1440 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1441 | lemma closed_translation: | 
| 69661 | 1442 | "closed ((+) a ` S)" if "closed S" for a :: "'a::real_normed_vector" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1443 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1444 |   have "(\<Union>x\<in> {a}. \<Union>y \<in> S. {x + y}) = ((+) a ` S)" by auto
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1445 | then show ?thesis | 
| 69661 | 1446 | using compact_closed_sums [OF compact_sing [of a] that] by auto | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1447 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1448 | |
| 69661 | 1449 | lemma closed_translation_subtract: | 
| 1450 | "closed ((\<lambda>x. x - a) ` S)" if "closed S" for a :: "'a::real_normed_vector" | |
| 1451 | using that closed_translation [of S "- a"] by (simp cong: image_cong_simp) | |
| 1452 | ||
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1453 | lemma closure_translation: | 
| 69661 | 1454 | "closure ((+) a ` s) = (+) a ` closure s" for a :: "'a::real_normed_vector" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1455 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1456 | have *: "(+) a ` (- s) = - (+) a ` s" | 
| 69661 | 1457 | by (auto intro!: image_eqI [where x = "x - a" for x]) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1458 | show ?thesis | 
| 69661 | 1459 | using interior_translation [of a "- s", symmetric] | 
| 1460 | by (simp add: closure_interior translation_Compl *) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1461 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1462 | |
| 69661 | 1463 | lemma closure_translation_subtract: | 
| 1464 | "closure ((\<lambda>x. x - a) ` s) = (\<lambda>x. x - a) ` closure s" for a :: "'a::real_normed_vector" | |
| 1465 | using closure_translation [of "- a" s] by (simp cong: image_cong_simp) | |
| 1466 | ||
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1467 | lemma frontier_translation: | 
| 69661 | 1468 | "frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector" | 
| 1469 | by (auto simp add: frontier_def translation_diff interior_translation closure_translation) | |
| 1470 | ||
| 1471 | lemma frontier_translation_subtract: | |
| 1472 | "frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector" | |
| 1473 | by (auto simp add: frontier_def translation_diff interior_translation closure_translation) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1474 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1475 | lemma sphere_translation: | 
| 69661 | 1476 | "sphere (a + c) r = (+) a ` sphere c r" for a :: "'n::real_normed_vector" | 
| 1477 | by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) | |
| 1478 | ||
| 1479 | lemma sphere_translation_subtract: | |
| 1480 | "sphere (c - a) r = (\<lambda>x. x - a) ` sphere c r" for a :: "'n::real_normed_vector" | |
| 1481 | using sphere_translation [of "- a" c] by (simp cong: image_cong_simp) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1482 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1483 | lemma cball_translation: | 
| 69661 | 1484 | "cball (a + c) r = (+) a ` cball c r" for a :: "'n::real_normed_vector" | 
| 1485 | by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) | |
| 1486 | ||
| 1487 | lemma cball_translation_subtract: | |
| 1488 | "cball (c - a) r = (\<lambda>x. x - a) ` cball c r" for a :: "'n::real_normed_vector" | |
| 1489 | using cball_translation [of "- a" c] by (simp cong: image_cong_simp) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1490 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1491 | lemma ball_translation: | 
| 69661 | 1492 | "ball (a + c) r = (+) a ` ball c r" for a :: "'n::real_normed_vector" | 
| 1493 | by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) | |
| 1494 | ||
| 1495 | lemma ball_translation_subtract: | |
| 1496 | "ball (c - a) r = (\<lambda>x. x - a) ` ball c r" for a :: "'n::real_normed_vector" | |
| 1497 | using ball_translation [of "- a" c] by (simp cong: image_cong_simp) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1498 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1499 | |
| 70136 | 1500 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Homeomorphisms\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1501 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1502 | lemma homeomorphic_scaling: | 
| 72643 | 1503 | fixes S :: "'a::real_normed_vector set" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1504 | assumes "c \<noteq> 0" | 
| 72643 | 1505 | shows "S homeomorphic ((\<lambda>x. c *\<^sub>R x) ` S)" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1506 | unfolding homeomorphic_minimal | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1507 | apply (rule_tac x="\<lambda>x. c *\<^sub>R x" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1508 | apply (rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI) | 
| 72643 | 1509 | using assms by (auto simp: continuous_intros) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1510 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1511 | lemma homeomorphic_translation: | 
| 72643 | 1512 | fixes S :: "'a::real_normed_vector set" | 
| 1513 | shows "S homeomorphic ((\<lambda>x. a + x) ` S)" | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1514 | unfolding homeomorphic_minimal | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1515 | apply (rule_tac x="\<lambda>x. a + x" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1516 | apply (rule_tac x="\<lambda>x. -a + x" in exI) | 
| 72643 | 1517 | by (auto simp: continuous_intros) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1518 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1519 | lemma homeomorphic_affinity: | 
| 72643 | 1520 | fixes S :: "'a::real_normed_vector set" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1521 | assumes "c \<noteq> 0" | 
| 72643 | 1522 | shows "S homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` S)" | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1523 | proof - | 
| 72643 | 1524 | have *: "(+) a ` (*\<^sub>R) c ` S = (\<lambda>x. a + c *\<^sub>R x) ` S" by auto | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1525 | show ?thesis | 
| 72643 | 1526 | by (metis "*" assms homeomorphic_scaling homeomorphic_trans homeomorphic_translation) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1527 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1528 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1529 | lemma homeomorphic_balls: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1530 | fixes a b ::"'a::real_normed_vector" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1531 | assumes "0 < d" "0 < e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1532 | shows "(ball a d) homeomorphic (ball b e)" (is ?th) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1533 | and "(cball a d) homeomorphic (cball b e)" (is ?cth) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1534 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1535 | show ?th unfolding homeomorphic_minimal | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1536 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1537 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1538 | using assms | 
| 72643 | 1539 | by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1540 | show ?cth unfolding homeomorphic_minimal | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1541 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1542 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1543 | using assms | 
| 72643 | 1544 | by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_le_eq) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1545 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1546 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1547 | lemma homeomorphic_spheres: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1548 | fixes a b ::"'a::real_normed_vector" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1549 | assumes "0 < d" "0 < e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1550 | shows "(sphere a d) homeomorphic (sphere b e)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1551 | unfolding homeomorphic_minimal | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1552 | apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1553 | apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1554 | using assms | 
| 72643 | 1555 | by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1556 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1557 | lemma homeomorphic_ball01_UNIV: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1558 | "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1559 | (is "?B homeomorphic ?U") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1560 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1561 | have "x \<in> (\<lambda>z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1562 | apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI) | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70724diff
changeset | 1563 | apply (auto simp: field_split_simps) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1564 | using norm_ge_zero [of x] apply linarith+ | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1565 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1566 | then show "(\<lambda>z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1567 | by blast | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1568 | have "x \<in> range (\<lambda>z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a | 
| 72643 | 1569 | using that | 
| 1570 | by (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI) (auto simp: field_split_simps) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1571 | then show "(\<lambda>z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70724diff
changeset | 1572 | by (force simp: field_split_simps dest: add_less_zeroD) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1573 | show "continuous_on (ball 0 1) (\<lambda>z. z /\<^sub>R (1 - norm z))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1574 | by (rule continuous_intros | force)+ | 
| 72643 | 1575 | have 0: "\<And>z. 1 + norm z \<noteq> 0" | 
| 1576 | by (metis (no_types) le_add_same_cancel1 norm_ge_zero not_one_le_zero) | |
| 1577 | then show "continuous_on UNIV (\<lambda>z. z /\<^sub>R (1 + norm z))" | |
| 1578 | by (auto intro!: continuous_intros) | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1579 | show "\<And>x. x \<in> ball 0 1 \<Longrightarrow> | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1580 | x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70724diff
changeset | 1581 | by (auto simp: field_split_simps) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1582 | show "\<And>y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" | 
| 72643 | 1583 | using 0 by (auto simp: field_split_simps) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1584 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1585 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1586 | proposition homeomorphic_ball_UNIV: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1587 | fixes a ::"'a::real_normed_vector" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1588 | assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1589 | using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1590 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1591 | |
| 70136 | 1592 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Discrete\<close> | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1593 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1594 | lemma finite_implies_discrete: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1595 | fixes S :: "'a::topological_space set" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1596 | assumes "finite (f ` S)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1597 | 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))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1598 | proof - | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1599 | 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 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1600 |   proof (cases "f ` S - {f x} = {}")
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1601 | case True | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1602 | with zero_less_numeral show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1603 | by (fastforce simp add: Set.image_subset_iff cong: conj_cong) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1604 | next | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1605 | case False | 
| 72643 | 1606 | then obtain z where "z \<in> S" "f z \<noteq> f x" | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1607 | by blast | 
| 72643 | 1608 |     moreover have finn: "finite {norm (z - f x) |z. z \<in> f ` S - {f x}}"
 | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1609 | using assms by simp | 
| 72643 | 1610 |     ultimately have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` S - {f x}}"
 | 
| 1611 | by (force intro: finite_imp_less_Inf) | |
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1612 | show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1613 | by (force intro!: * cInf_le_finite [OF finn]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1614 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1615 | with assms show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1616 | by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1617 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1618 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1619 | |
| 70136 | 1620 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Completeness of "Isometry" (up to constant bounds)\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1621 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1622 | lemma cauchy_isometric:\<comment> \<open>TODO: rename lemma to \<open>Cauchy_isometric\<close>\<close> | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1623 | assumes e: "e > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1624 | and s: "subspace s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1625 | and f: "bounded_linear f" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1626 | and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1627 | and xs: "\<forall>n. x n \<in> s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1628 | and cf: "Cauchy (f \<circ> x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1629 | shows "Cauchy x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1630 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1631 | interpret f: bounded_linear f by fact | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1632 | have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" if "d > 0" for d :: real | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1633 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1634 | from that obtain N where N: "\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1635 | using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1636 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1637 | have "norm (x n - x N) < d" if "n \<ge> N" for n | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1638 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1639 | have "e * norm (x n - x N) \<le> norm (f (x n - x N))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1640 | using subspace_diff[OF s, of "x n" "x N"] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1641 | using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1642 | using normf[THEN bspec[where x="x n - x N"]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1643 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1644 | also have "norm (f (x n - x N)) < e * d" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1645 | using \<open>N \<le> n\<close> N unfolding f.diff[symmetric] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1646 | finally show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1647 | using \<open>e>0\<close> by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1648 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1649 | then show ?thesis by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1650 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1651 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1652 | by (simp add: Cauchy_altdef2 dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1653 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1654 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1655 | lemma complete_isometric_image: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1656 | assumes "0 < e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1657 | and s: "subspace s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1658 | and f: "bounded_linear f" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1659 | and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1660 | and cs: "complete s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1661 | shows "complete (f ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1662 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1663 | have "\<exists>l\<in>f ` s. (g \<longlongrightarrow> l) sequentially" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1664 | if as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g" for g | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1665 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1666 | from that obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1667 | using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1668 | then have x: "\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)" by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1669 | then have "f \<circ> x = g" by (simp add: fun_eq_iff) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1670 | then obtain l where "l\<in>s" and l:"(x \<longlongrightarrow> l) sequentially" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1671 | using cs[unfolded complete_def, THEN spec[where x=x]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1672 | using cauchy_isometric[OF \<open>0 < e\<close> s f normf] and cfg and x(1) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1673 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1674 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1675 | using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1676 | by (auto simp: \<open>f \<circ> x = g\<close>) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1677 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1678 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1679 | unfolding complete_def by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1680 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1681 | |
| 69617 | 1682 | subsection \<open>Connected Normed Spaces\<close> | 
| 1683 | ||
| 1684 | lemma compact_components: | |
| 1685 | fixes s :: "'a::heine_borel set" | |
| 1686 | shows "\<lbrakk>compact s; c \<in> components s\<rbrakk> \<Longrightarrow> compact c" | |
| 1687 | by (meson bounded_subset closed_components in_components_subset compact_eq_bounded_closed) | |
| 1688 | ||
| 1689 | lemma discrete_subset_disconnected: | |
| 1690 | fixes S :: "'a::topological_space set" | |
| 1691 | fixes t :: "'b::real_normed_vector set" | |
| 1692 | assumes conf: "continuous_on S f" | |
| 1693 | 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)" | |
| 1694 |    shows "f ` S \<subseteq> {y. connected_component_set (f ` S) y = {y}}"
 | |
| 1695 | proof - | |
| 1696 |   { fix x assume x: "x \<in> S"
 | |
| 1697 | 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)" | |
| 1698 | using conf no [OF x] by auto | |
| 72643 | 1699 | then have e2: "0 \<le> e/2" | 
| 69617 | 1700 | by simp | 
| 72643 | 1701 | define F where "F \<equiv> connected_component_set (f ` S) (f x)" | 
| 1702 | have False if "y \<in> S" and ccs: "f y \<in> F" and not: "f y \<noteq> f x" for y | |
| 1703 | proof - | |
| 1704 | define C where "C \<equiv> cball (f x) (e/2)" | |
| 1705 | define D where "D \<equiv> - ball (f x) e" | |
| 1706 |       have disj: "C \<inter> D = {}"
 | |
| 1707 | unfolding C_def D_def using \<open>0 < e\<close> by fastforce | |
| 1708 | moreover have FCD: "F \<subseteq> C \<union> D" | |
| 1709 | proof - | |
| 1710 | have "t \<in> C \<or> t \<in> D" if "t \<in> F" for t | |
| 1711 | proof - | |
| 1712 | obtain y where "y \<in> S" "t = f y" | |
| 1713 | using F_def \<open>t \<in> F\<close> connected_component_in by blast | |
| 1714 | then show ?thesis | |
| 1715 | by (metis C_def ComplI D_def centre_in_cball dist_norm e2 ele mem_ball norm_minus_commute not_le) | |
| 1716 | qed | |
| 1717 | then show ?thesis | |
| 1718 | by auto | |
| 1719 | qed | |
| 1720 |       ultimately have "C \<inter> F = {} \<or> D \<inter> F = {}"
 | |
| 1721 | using connected_closed [of "F"] \<open>e>0\<close> not | |
| 1722 | unfolding C_def D_def | |
| 1723 | by (metis Elementary_Metric_Spaces.open_ball F_def closed_cball connected_connected_component inf_bot_left open_closed) | |
| 1724 |       moreover have "C \<inter> F \<noteq> {}"
 | |
| 1725 | unfolding disjoint_iff | |
| 1726 | by (metis FCD ComplD image_eqI mem_Collect_eq subsetD x D_def F_def Un_iff \<open>0 < e\<close> centre_in_ball connected_component_refl_eq) | |
| 1727 |       moreover have "D \<inter> F \<noteq> {}"
 | |
| 1728 | unfolding disjoint_iff | |
| 1729 | by (metis ComplI D_def ccs dist_norm ele mem_ball norm_minus_commute not not_le that(1)) | |
| 1730 | ultimately show ?thesis by metis | |
| 1731 | qed | |
| 69617 | 1732 | moreover have "connected_component_set (f ` S) (f x) \<subseteq> f ` S" | 
| 1733 | by (auto simp: connected_component_in) | |
| 1734 |     ultimately have "connected_component_set (f ` S) (f x) = {f x}"
 | |
| 72643 | 1735 | by (auto simp: x F_def) | 
| 69617 | 1736 | } | 
| 1737 | with assms show ?thesis | |
| 1738 | by blast | |
| 1739 | qed | |
| 1740 | ||
| 1741 | lemma continuous_disconnected_range_constant_eq: | |
| 1742 | "(connected S \<longleftrightarrow> | |
| 1743 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | |
| 1744 |             \<forall>t. continuous_on S f \<and> f ` S \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
 | |
| 1745 | \<longrightarrow> f constant_on S))" (is ?thesis1) | |
| 1746 | and continuous_discrete_range_constant_eq: | |
| 1747 | "(connected S \<longleftrightarrow> | |
| 1748 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | |
| 1749 | continuous_on S f \<and> | |
| 1750 | (\<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))) | |
| 1751 | \<longrightarrow> f constant_on S))" (is ?thesis2) | |
| 1752 | and continuous_finite_range_constant_eq: | |
| 1753 | "(connected S \<longleftrightarrow> | |
| 1754 | (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1. | |
| 1755 | continuous_on S f \<and> finite (f ` S) | |
| 1756 | \<longrightarrow> f constant_on S))" (is ?thesis3) | |
| 1757 | proof - | |
| 1758 | have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk> | |
| 1759 | \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)" | |
| 1760 | by blast | |
| 1761 | have "?thesis1 \<and> ?thesis2 \<and> ?thesis3" | |
| 1762 | apply (rule *) | |
| 1763 | using continuous_disconnected_range_constant apply metis | |
| 1764 | apply clarify | |
| 1765 | apply (frule discrete_subset_disconnected; blast) | |
| 1766 | apply (blast dest: finite_implies_discrete) | |
| 1767 | apply (blast intro!: finite_range_constant_imp_connected) | |
| 1768 | done | |
| 1769 | then show ?thesis1 ?thesis2 ?thesis3 | |
| 1770 | by blast+ | |
| 1771 | qed | |
| 1772 | ||
| 1773 | lemma continuous_discrete_range_constant: | |
| 1774 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" | |
| 1775 | assumes S: "connected S" | |
| 1776 | and "continuous_on S f" | |
| 1777 | 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)" | |
| 1778 | shows "f constant_on S" | |
| 1779 | using continuous_discrete_range_constant_eq [THEN iffD1, OF S] assms by blast | |
| 1780 | ||
| 1781 | lemma continuous_finite_range_constant: | |
| 1782 | fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1" | |
| 1783 | assumes "connected S" | |
| 1784 | and "continuous_on S f" | |
| 1785 | and "finite (f ` S)" | |
| 1786 | shows "f constant_on S" | |
| 1787 | using assms continuous_finite_range_constant_eq by blast | |
| 1788 | ||
| 70630 | 1789 | end |