| author | wenzelm | 
| Fri, 22 Sep 2023 16:12:10 +0200 | |
| changeset 78684 | aa532cf1c894 | 
| parent 78131 | 1cadc477f644 | 
| permissions | -rw-r--r-- | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 1 | (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 2 | License: BSD | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 3 | *) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 4 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 5 | section \<open>Metrics on product spaces\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 6 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 7 | theory Function_Metric | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 8 | imports | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 9 | Function_Topology | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 10 | Elementary_Metric_Spaces | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 11 | begin | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 12 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 13 | text \<open>In general, the product topology is not metrizable, unless the index set is countable. | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 14 | When the index set is countable, essentially any (convergent) combination of the metrics on the | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 15 | factors will do. We use below the simplest one, based on \<open>L\<^sup>1\<close>, but \<open>L\<^sup>2\<close> would also work, | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 16 | for instance. | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 17 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 18 | What is not completely trivial is that the distance thus defined induces the same topology | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 19 | as the product topology. This is what we have to prove to show that we have an instance | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 20 | of \<^class>\<open>metric_space\<close>. | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 21 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 22 | The proofs below would work verbatim for general countable products of metric spaces. However, | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 23 | since distances are only implemented in terms of type classes, we only develop the theory | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 24 | for countable products of the same space.\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 25 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 26 | instantiation "fun" :: (countable, metric_space) metric_space | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 27 | begin | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 28 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 29 | definition\<^marker>\<open>tag important\<close> dist_fun_def: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 30 | "dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 31 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 32 | definition\<^marker>\<open>tag important\<close> uniformity_fun_def: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 33 |   "(uniformity::(('a \<Rightarrow> 'b) \<times> ('a \<Rightarrow> 'b)) filter) = (INF e\<in>{0<..}. principal {(x, y). dist (x::('a\<Rightarrow>'b)) y < e})"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 34 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 35 | text \<open>Except for the first one, the auxiliary lemmas below are only useful when proving the | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 36 | instance: once it is proved, they become trivial consequences of the general theory of metric | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 37 | spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 38 | to do this.\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 39 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 40 | lemma dist_fun_le_dist_first_terms: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 41 |   "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 42 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 43 | have "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 44 | = (\<Sum>n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 45 | by (rule suminf_cong, simp add: power_add) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 46 | also have "... = (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 47 | apply (rule suminf_mult) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 48 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 49 | also have "... \<le> (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 50 | apply (simp, rule suminf_le, simp) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 51 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 52 | also have "... = (1/2)^(Suc N) * 2" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 53 | using suminf_geometric[of "1/2"] by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 54 | also have "... = (1/2)^N" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 55 | by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 56 | finally have *: "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) \<le> (1/2)^N" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 57 | by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 58 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 59 |   define M where "M = Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 60 | have "dist (x (from_nat 0)) (y (from_nat 0)) \<le> M" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 61 | unfolding M_def by (rule Max_ge, auto) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 62 | then have [simp]: "M \<ge> 0" by (meson dual_order.trans zero_le_dist) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 63 | have "dist (x (from_nat n)) (y (from_nat n)) \<le> M" if "n\<le>N" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 64 | unfolding M_def apply (rule Max_ge) using that by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 65 | then have i: "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le> M" if "n\<le>N" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 66 | using that by force | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 67 | have "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 68 | (\<Sum>n< Suc N. M * (1 / 2) ^ n)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 69 | by (rule sum_mono, simp add: i) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 70 | also have "... = M * (\<Sum>n<Suc N. (1 / 2) ^ n)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 71 | by (rule sum_distrib_left[symmetric]) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 72 | also have "... \<le> M * (\<Sum>n. (1 / 2) ^ n)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 73 | by (rule mult_left_mono, rule sum_le_suminf, auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 74 | also have "... = M * 2" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 75 | using suminf_geometric[of "1/2"] by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 76 | finally have **: "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le> 2 * M" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 77 | by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 78 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 79 | have "dist x y = (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 80 | unfolding dist_fun_def by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 81 | also have "... = (\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 82 | + (\<Sum>n<Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 83 | apply (rule suminf_split_initial_segment) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 84 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 85 | also have "... \<le> 2 * M + (1/2)^N" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 86 | using * ** by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 87 | finally show ?thesis unfolding M_def by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 88 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 89 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 90 | lemma open_fun_contains_ball_aux: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 91 |   assumes "open (U::(('a \<Rightarrow> 'b) set))"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 92 | "x \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 93 |   shows "\<exists>e>0. {y. dist x y < e} \<subseteq> U"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 94 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 95 | have *: "openin (product_topology (\<lambda>i. euclidean) UNIV) U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 96 | using open_fun_def assms by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 97 | obtain X where H: "Pi\<^sub>E UNIV X \<subseteq> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 98 | "\<And>i. openin euclidean (X i)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 99 |                     "finite {i. X i \<noteq> topspace euclidean}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 100 | "x \<in> Pi\<^sub>E UNIV X" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 101 | using product_topology_open_contains_basis[OF * \<open>x \<in> U\<close>] by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 102 |   define I where "I = {i. X i \<noteq> topspace euclidean}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 103 | have "finite I" unfolding I_def using H(3) by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 104 |   {
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 105 | fix i | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 106 | have "x i \<in> X i" using \<open>x \<in> U\<close> H by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 107 | then have "\<exists>e. e>0 \<and> ball (x i) e \<subseteq> X i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 108 | using \<open>openin euclidean (X i)\<close> open_openin open_contains_ball by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 109 | then obtain e where "e>0" "ball (x i) e \<subseteq> X i" by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 110 | define f where "f = min e (1/2)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 111 | have "f>0" "f<1" unfolding f_def using \<open>e>0\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 112 | moreover have "ball (x i) f \<subseteq> X i" unfolding f_def using \<open>ball (x i) e \<subseteq> X i\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 113 | ultimately have "\<exists>f. f > 0 \<and> f < 1 \<and> ball (x i) f \<subseteq> X i" by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 114 | } note * = this | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 115 | have "\<exists>e. \<forall>i. e i > 0 \<and> e i < 1 \<and> ball (x i) (e i) \<subseteq> X i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 116 | by (rule choice, auto simp add: *) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 117 | then obtain e where "\<And>i. e i > 0" "\<And>i. e i < 1" "\<And>i. ball (x i) (e i) \<subseteq> X i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 118 | by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 119 |   define m where "m = Min {(1/2)^(to_nat i) * e i|i. i \<in> I}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 120 |   have "m > 0" if "I\<noteq>{}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 121 |     unfolding m_def Min_gr_iff using \<open>finite I\<close> \<open>I \<noteq> {}\<close> \<open>\<And>i. e i > 0\<close> by auto
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 122 |   moreover have "{y. dist x y < m} \<subseteq> U"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 123 | proof (auto) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 124 | fix y assume "dist x y < m" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 125 | have "y i \<in> X i" if "i \<in> I" for i | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 126 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 127 | have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 128 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 129 | define n where "n = to_nat i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 130 | have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 < m" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 131 |         using \<open>dist x y < m\<close> unfolding dist_fun_def using sum_le_suminf[OF *, of "{n}"] by auto
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 132 | then have "(1/2)^(to_nat i) * min (dist (x i) (y i)) 1 < m" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 133 | using \<open>n = to_nat i\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 134 | also have "... \<le> (1/2)^(to_nat i) * e i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 135 | unfolding m_def apply (rule Min_le) using \<open>finite I\<close> \<open>i \<in> I\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 136 | ultimately have "min (dist (x i) (y i)) 1 < e i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 137 | by (auto simp add: field_split_simps) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 138 | then have "dist (x i) (y i) < e i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 139 | using \<open>e i < 1\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 140 | then show "y i \<in> X i" using \<open>ball (x i) (e i) \<subseteq> X i\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 141 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 142 | then have "y \<in> Pi\<^sub>E UNIV X" using H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 143 | then show "y \<in> U" using \<open>Pi\<^sub>E UNIV X \<subseteq> U\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 144 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 145 |   ultimately have *: "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I \<noteq> {}" using that by auto
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 146 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 147 |   have "U = UNIV" if "I = {}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 148 | using that H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 149 |   then have "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I = {}" using that zero_less_one by blast
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 150 |   then show "\<exists>m>0. {y. dist x y < m} \<subseteq> U" using * by blast
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 151 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 152 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 153 | lemma ball_fun_contains_open_aux: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 154 |   fixes x::"('a \<Rightarrow> 'b)" and e::real
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 155 | assumes "e>0" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 156 |   shows "\<exists>U. open U \<and> x \<in> U \<and> U \<subseteq> {y. dist x y < e}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 157 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 158 | have "\<exists>N::nat. 2^N > 8/e" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 159 | by (simp add: real_arch_pow) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 160 | then obtain N::nat where "2^N > 8/e" by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 161 | define f where "f = e/4" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 162 | have [simp]: "e>0" "f > 0" unfolding f_def using assms by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 163 |   define X::"('a \<Rightarrow> 'b set)" where "X = (\<lambda>i. if (\<exists>n\<le>N. i = from_nat n) then {z. dist (x i) z < f} else UNIV)"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 164 |   have "{i. X i \<noteq> UNIV} \<subseteq> from_nat`{0..N}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 165 | unfolding X_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 166 |   then have "finite {i. X i \<noteq> topspace euclidean}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 167 | unfolding topspace_euclidean using finite_surj by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 168 | have "\<And>i. open (X i)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 169 | unfolding X_def using metric_space_class.open_ball open_UNIV by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 170 | then have "\<And>i. openin euclidean (X i)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 171 | using open_openin by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 172 | define U where "U = Pi\<^sub>E UNIV X" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 173 | have "open U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 174 | unfolding open_fun_def product_topology_def apply (rule topology_generated_by_Basis) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 175 |     unfolding U_def using \<open>\<And>i. openin euclidean (X i)\<close> \<open>finite {i. X i \<noteq> topspace euclidean}\<close>
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 176 | by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 177 | moreover have "x \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 178 | unfolding U_def X_def by (auto simp add: PiE_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 179 | moreover have "dist x y < e" if "y \<in> U" for y | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 180 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 181 | have *: "dist (x (from_nat n)) (y (from_nat n)) \<le> f" if "n \<le> N" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 182 | using \<open>y \<in> U\<close> unfolding U_def apply (auto simp add: PiE_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 183 | unfolding X_def using that by (metis less_imp_le mem_Collect_eq) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 184 |     have **: "Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} \<le> f"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 185 | apply (rule Max.boundedI) using * by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 186 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 187 |     have "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 188 | by (rule dist_fun_le_dist_first_terms) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 189 | also have "... \<le> 2 * f + e / 8" | 
| 71633 | 190 | apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: field_split_simps) | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 191 | also have "... \<le> e/2 + e/8" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 192 | unfolding f_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 193 | also have "... < e" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 194 | by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 195 | finally show "dist x y < e" by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 196 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 197 | ultimately show ?thesis by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 198 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 199 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 200 | lemma fun_open_ball_aux: | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 201 |   fixes U::"('a \<Rightarrow> 'b) set"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 202 | shows "open U \<longleftrightarrow> (\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 203 | proof (auto) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 204 | assume "open U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 205 | fix x assume "x \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 206 | then show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 207 | using open_fun_contains_ball_aux[OF \<open>open U\<close> \<open>x \<in> U\<close>] by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 208 | next | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 209 | assume H: "\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 210 |   define K where "K = {V. open V \<and> V \<subseteq> U}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 211 |   {
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 212 | fix x assume "x \<in> U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 213 |     then obtain e where e: "e>0" "{y. dist x y < e} \<subseteq> U" using H by blast
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 214 |     then obtain V where V: "open V" "x \<in> V" "V \<subseteq> {y. dist x y < e}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 215 | using ball_fun_contains_open_aux[OF \<open>e>0\<close>, of x] by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 216 | have "V \<in> K" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 217 | unfolding K_def using e(2) V(1) V(3) by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 218 | then have "x \<in> (\<Union>K)" using \<open>x \<in> V\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 219 | } | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 220 | then have "(\<Union>K) = U" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 221 | unfolding K_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 222 | moreover have "open (\<Union>K)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 223 | unfolding K_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 224 | ultimately show "open U" by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 225 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 226 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 227 | instance proof | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 228 | fix x y::"'a \<Rightarrow> 'b" show "(dist x y = 0) = (x = y)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 229 | proof | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 230 | assume "x = y" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 231 | then show "dist x y = 0" unfolding dist_fun_def using \<open>x = y\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 232 | next | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 233 | assume "dist x y = 0" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 234 | have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 235 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 236 | have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 = 0" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 237 | using \<open>dist x y = 0\<close> unfolding dist_fun_def by (simp add: "*" suminf_eq_zero_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 238 | then have "dist (x (from_nat n)) (y (from_nat n)) = 0" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 239 | by (metis div_0 min_def nonzero_mult_div_cancel_left power_eq_0_iff | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 240 | zero_eq_1_divide_iff zero_neq_numeral) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 241 | then have "x (from_nat n) = y (from_nat n)" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 242 | by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 243 | then have "x i = y i" for i | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 244 | by (metis from_nat_to_nat) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 245 | then show "x = y" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 246 | by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 247 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 248 | next | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 249 | text\<open>The proof of the triangular inequality is trivial, modulo the fact that we are dealing | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 250 | with infinite series, hence we should justify the convergence at each step. In this | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 251 | respect, the following simplification rule is extremely handy.\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 252 | have [simp]: "summable (\<lambda>n. (1/2)^n * min (dist (u (from_nat n)) (v (from_nat n))) 1)" for u v::"'a \<Rightarrow> 'b" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 253 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 254 | fix x y z::"'a \<Rightarrow> 'b" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 255 |   {
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 256 | fix n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 257 | have *: "dist (x (from_nat n)) (y (from_nat n)) \<le> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 258 | dist (x (from_nat n)) (z (from_nat n)) + dist (y (from_nat n)) (z (from_nat n))" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 259 | by (simp add: dist_triangle2) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 260 | have "0 \<le> dist (y (from_nat n)) (z (from_nat n))" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 261 | using zero_le_dist by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 262 | moreover | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 263 |     {
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 264 | assume "min (dist (y (from_nat n)) (z (from_nat n))) 1 \<noteq> dist (y (from_nat n)) (z (from_nat n))" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 265 | then have "1 \<le> min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 266 | by (metis (no_types) diff_le_eq diff_self min_def zero_le_dist zero_le_one) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 267 | } | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 268 | ultimately have "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 269 | min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 270 | using * by linarith | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 271 | } note ineq = this | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 272 | have "dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 273 | unfolding dist_fun_def by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 274 | also have "... \<le> (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 275 | + (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 276 | apply (rule suminf_le) | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
71633diff
changeset | 277 | using ineq apply (metis (no_types, opaque_lifting) add.right_neutral distrib_left | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 278 | le_divide_eq_numeral1(1) mult_2_right mult_left_mono zero_le_one zero_le_power) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 279 | by (auto simp add: summable_add) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 280 | also have "... = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 281 | + (\<Sum>n. (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 282 | by (rule suminf_add[symmetric], auto) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 283 | also have "... = dist x z + dist y z" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 284 | unfolding dist_fun_def by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 285 | finally show "dist x y \<le> dist x z + dist y z" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 286 | by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 287 | next | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 288 | text\<open>Finally, we show that the topology generated by the distance and the product | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 289 | topology coincide. This is essentially contained in Lemma \<open>fun_open_ball_aux\<close>, | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 290 | except that the condition to prove is expressed with filters. To deal with this, | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 291 | we copy some mumbo jumbo from Lemma \<open>eventually_uniformity_metric\<close> in | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 292 | \<^file>\<open>~~/src/HOL/Real_Vector_Spaces.thy\<close>\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 293 |   fix U::"('a \<Rightarrow> 'b) set"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 294 |   have "eventually P uniformity \<longleftrightarrow> (\<exists>e>0. \<forall>x (y::('a \<Rightarrow> 'b)). dist x y < e \<longrightarrow> P (x, y))" for P
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 295 | unfolding uniformity_fun_def apply (subst eventually_INF_base) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 296 | by (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 297 | then show "open U = (\<forall>x\<in>U. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> U)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 298 | unfolding fun_open_ball_aux by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 299 | qed (simp add: uniformity_fun_def) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 300 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 301 | end | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 302 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 303 | text \<open>Nice properties of spaces are preserved under countable products. In addition | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 304 | to first countability, second countability and metrizability, as we have seen above, | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 305 | completeness is also preserved, and therefore being Polish.\<close> | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 306 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 307 | instance "fun" :: (countable, complete_space) complete_space | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 308 | proof | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 309 |   fix u::"nat \<Rightarrow> ('a \<Rightarrow> 'b)" assume "Cauchy u"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 310 | have "Cauchy (\<lambda>n. u n i)" for i | 
| 78131 
1cadc477f644
Even more material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 311 | unfolding Cauchy_def | 
| 
1cadc477f644
Even more material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 312 | proof (intro strip) | 
| 
1cadc477f644
Even more material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 313 | fix e::real assume "e>0" | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 314 | obtain k where "i = from_nat k" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 315 | using from_nat_to_nat[of i] by metis | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 316 | have "(1/2)^k * min e 1 > 0" using \<open>e>0\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 317 | then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1" | 
| 78131 
1cadc477f644
Even more material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 318 | using \<open>Cauchy u\<close> by (meson Cauchy_def) | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 319 | then obtain N::nat where C: "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 320 | by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 321 | have "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 322 | proof (auto) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 323 | fix m n::nat assume "m \<ge> N" "n \<ge> N" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 324 | have "(1/2)^k * min (dist (u m i) (u n i)) 1 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 325 |               = sum (\<lambda>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1) {k}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 326 | using \<open>i = from_nat k\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 327 | also have "... \<le> (\<Sum>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 328 | apply (rule sum_le_suminf) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 329 | by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 330 | also have "... = dist (u m) (u n)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 331 | unfolding dist_fun_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 332 | also have "... < (1/2)^k * min e 1" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 333 | using C \<open>m \<ge> N\<close> \<open>n \<ge> N\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 334 | finally have "min (dist (u m i) (u n i)) 1 < min e 1" | 
| 71633 | 335 | by (auto simp add: field_split_simps) | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 336 | then show "dist (u m i) (u n i) < e" by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 337 | qed | 
| 78131 
1cadc477f644
Even more material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 338 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (u m i) (u n i) < e" | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 339 | by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 340 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 341 | then have "\<exists>x. (\<lambda>n. u n i) \<longlonglongrightarrow> x" for i | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 342 | using Cauchy_convergent convergent_def by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 343 | then have "\<exists>x. \<forall>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 344 | using choice by force | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 345 | then obtain x where *: "\<And>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i" by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 346 | have "u \<longlonglongrightarrow> x" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 347 | proof (rule metric_LIMSEQ_I) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 348 | fix e assume [simp]: "e>(0::real)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 349 | have i: "\<exists>K. \<forall>n\<ge>K. dist (u n i) (x i) < e/4" for i | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 350 | by (rule metric_LIMSEQ_D, auto simp add: *) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 351 | have "\<exists>K. \<forall>i. \<forall>n\<ge>K i. dist (u n i) (x i) < e/4" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 352 | apply (rule choice) using i by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 353 | then obtain K where K: "\<And>i n. n \<ge> K i \<Longrightarrow> dist (u n i) (x i) < e/4" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 354 | by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 355 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 356 | have "\<exists>N::nat. 2^N > 4/e" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 357 | by (simp add: real_arch_pow) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 358 | then obtain N::nat where "2^N > 4/e" by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 359 |     define L where "L = Max {K (from_nat n)|n. n \<le> N}"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 360 | have "dist (u k) x < e" if "k \<ge> L" for k | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 361 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 362 | have *: "dist (u k (from_nat n)) (x (from_nat n)) \<le> e / 4" if "n \<le> N" for n | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 363 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 364 | have "K (from_nat n) \<le> L" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 365 | unfolding L_def apply (rule Max_ge) using \<open>n \<le> N\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 366 | then have "k \<ge> K (from_nat n)" using \<open>k \<ge> L\<close> by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 367 | then show ?thesis using K less_imp_le by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 368 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 369 |       have **: "Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} \<le> e/4"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 370 | apply (rule Max.boundedI) using * by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 371 |       have "dist (u k) x \<le> 2 * Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} + (1/2)^N"
 | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 372 | using dist_fun_le_dist_first_terms by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 373 | also have "... \<le> 2 * e/4 + e/4" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 374 | apply (rule add_mono) | 
| 71633 | 375 | using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: field_split_simps) | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 376 | also have "... < e" by auto | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 377 | finally show ?thesis by simp | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 378 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 379 | then show "\<exists>L. \<forall>k\<ge>L. dist (u k) x < e" by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 380 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 381 | then show "convergent u" using convergent_def by blast | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 382 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 383 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 384 | instance "fun" :: (countable, polish_space) polish_space | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 385 | by standard | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 386 | |
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: diff
changeset | 387 | end |