| author | paulson <lp15@cam.ac.uk> | 
| Tue, 25 Jan 2022 14:13:33 +0000 | |
| changeset 75008 | 43b3d5318d72 | 
| parent 71633 | 07bec530f02e | 
| child 76821 | 337c2265d8a2 | 
| permissions | -rw-r--r-- | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 1 | (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 2 | License: BSD | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 3 | *) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 4 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 5 | theory Function_Topology | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 6 | imports | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 7 | Elementary_Topology | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 8 | Abstract_Limits | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 9 | Connected | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 10 | begin | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 11 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 12 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 13 | section \<open>Function Topology\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 14 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 15 | text \<open>We want to define the general product topology. | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 16 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 17 | The product topology on a product of topological spaces is generated by | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 18 | the sets which are products of open sets along finitely many coordinates, and the whole | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 19 | space along the other coordinates. This is the coarsest topology for which the projection | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 20 | to each factor is continuous. | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 21 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 22 | To form a product of objects in Isabelle/HOL, all these objects should be subsets of a common type | 
| 69597 | 23 | 'a. The product is then \<^term>\<open>Pi\<^sub>E I X\<close>, the set of elements from \<open>'i\<close> to \<open>'a\<close> such that the \<open>i\<close>-th | 
| 69566 | 24 | coordinate belongs to \<open>X i\<close> for all \<open>i \<in> I\<close>. | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 25 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 26 | Hence, to form a product of topological spaces, all these spaces should be subsets of a common type. | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 27 | This means that type classes can not be used to define such a product if one wants to take the | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 28 | product of different topological spaces (as the type 'a can only be given one structure of | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 29 | topological space using type classes). On the other hand, one can define different topologies (as | 
| 69566 | 30 | introduced in \<open>thy\<close>) on one type, and these topologies do not need to | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 31 | share the same maximal open set. Hence, one can form a product of topologies in this sense, and | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 32 | this works well. The big caveat is that it does not interact well with the main body of | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 33 | topology in Isabelle/HOL defined in terms of type classes... For instance, continuity of maps | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 34 | is not defined in this setting. | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 35 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 36 | As the product of different topological spaces is very important in several areas of | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 37 | mathematics (for instance adeles), I introduce below the product topology in terms of topologies, | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 38 | and reformulate afterwards the consequences in terms of type classes (which are of course very | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 39 | handy for applications). | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 40 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 41 | Given this limitation, it looks to me that it would be very beneficial to revamp the theory | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 42 | of topological spaces in Isabelle/HOL in terms of topologies, and keep the statements involving | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 43 | type classes as consequences of more general statements in terms of topologies (but I am | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 44 | probably too naive here). | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 45 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 46 | Here is an example of a reformulation using topologies. Let | 
| 69566 | 47 | @{text [display]
 | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 48 | \<open>continuous_map T1 T2 f = | 
| 69566 | 49 | ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) | 
| 50 | \<and> (f`(topspace T1) \<subseteq> (topspace T2)))\<close>} | |
| 51 | be the natural continuity definition of a map from the topology \<open>T1\<close> to the topology \<open>T2\<close>. Then | |
| 52 | the current \<open>continuous_on\<close> (with type classes) can be redefined as | |
| 53 | @{text [display] \<open>continuous_on s f =
 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 54 | continuous_map (top_of_set s) (topology euclidean) f\<close>} | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 55 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 56 | In fact, I need \<open>continuous_map\<close> to express the continuity of the projection on subfactors | 
| 69566 | 57 | for the product topology, in Lemma~\<open>continuous_on_restrict_product_topology\<close>, and I show | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 58 | the above equivalence in Lemma~\<open>continuous_map_iff_continuous\<close>. | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 59 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 60 | I only develop the basics of the product topology in this theory. The most important missing piece | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 61 | is Tychonov theorem, stating that a product of compact spaces is always compact for the product | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 62 | topology, even when the product is not finite (or even countable). | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 63 | |
| 69566 | 64 | I realized afterwards that this theory has a lot in common with \<^file>\<open>~~/src/HOL/Library/Finite_Map.thy\<close>. | 
| 64911 | 65 | \<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 66 | |
| 69683 | 67 | subsection \<open>The product topology\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 68 | |
| 64911 | 69 | text \<open>We can now define the product topology, as generated by | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 70 | the sets which are products of open sets along finitely many coordinates, and the whole | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 71 | space along the other coordinates. Equivalently, it is generated by sets which are one open | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 72 | set along one single coordinate, and the whole space along other coordinates. In fact, this is only | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 73 | equivalent for nonempty products, but for the empty product the first formulation is better | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 74 | (the second one gives an empty product space, while an empty product should have exactly one | 
| 69566 | 75 | point, equal to \<open>undefined\<close> along all coordinates. | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 76 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 77 | So, we use the first formulation, which moreover seems to give rise to more straightforward proofs. | 
| 64911 | 78 | \<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 79 | |
| 70136 | 80 | definition\<^marker>\<open>tag important\<close> product_topology::"('i \<Rightarrow> ('a topology)) \<Rightarrow> ('i set) \<Rightarrow> (('i \<Rightarrow> 'a) topology)"
 | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 81 | where "product_topology T I = | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 82 |     topology_generated_by {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 83 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 84 | abbreviation powertop_real :: "'a set \<Rightarrow> ('a \<Rightarrow> real) topology"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 85 | where "powertop_real \<equiv> product_topology (\<lambda>i. euclideanreal)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 86 | |
| 64911 | 87 | text \<open>The total set of the product topology is the product of the total sets | 
| 88 | along each coordinate.\<close> | |
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 89 | |
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 90 | proposition product_topology: | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 91 | "product_topology X I = | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 92 | topology | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 93 | (arbitrary union_of | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 94 | ((finite intersection_of | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 95 |         (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U))
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 96 | relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))))" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 97 | (is "_ = topology (_ union_of ((_ intersection_of ?\<Psi>) relative_to ?TOP))") | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 98 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 99 |   let ?\<Omega> = "(\<lambda>F. \<exists>Y. F = Pi\<^sub>E I Y \<and> (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)})"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 100 | have *: "(finite' intersection_of ?\<Omega>) A = (finite intersection_of ?\<Psi> relative_to ?TOP) A" for A | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 101 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 102 | have 1: "\<exists>U. (\<exists>\<U>. finite \<U> \<and> \<U> \<subseteq> Collect ?\<Psi> \<and> \<Inter>\<U> = U) \<and> ?TOP \<inter> U = \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 103 |       if \<U>: "\<U> \<subseteq> Collect ?\<Omega>" and "finite' \<U>" "A = \<Inter>\<U>" "\<U> \<noteq> {}" for \<U>
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 104 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 105 |       have "\<forall>U \<in> \<U>. \<exists>Y. U = Pi\<^sub>E I Y \<and> (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 106 | using \<U> by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 107 |       then obtain Y where Y: "\<And>U. U \<in> \<U> \<Longrightarrow> U = Pi\<^sub>E I (Y U) \<and> (\<forall>i. openin (X i) (Y U i)) \<and> finite {i. (Y U) i \<noteq> topspace (X i)}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 108 | by metis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 109 | obtain U where "U \<in> \<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 110 |         using \<open>\<U> \<noteq> {}\<close> by blast
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 111 |       let ?F = "\<lambda>U. (\<lambda>i. {f. f i \<in> Y U i}) ` {i \<in> I. Y U i \<noteq> topspace (X i)}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 112 | show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 113 | proof (intro conjI exI) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 114 | show "finite (\<Union>U\<in>\<U>. ?F U)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 115 | using Y \<open>finite' \<U>\<close> by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 116 | show "?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U) = \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 117 | proof | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 118 | have *: "f \<in> U" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 119 | if "U \<in> \<U>" and "\<forall>V\<in>\<U>. \<forall>i. i \<in> I \<and> Y V i \<noteq> topspace (X i) \<longrightarrow> f i \<in> Y V i" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 120 | and "\<forall>i\<in>I. f i \<in> topspace (X i)" and "f \<in> extensional I" for f U | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 121 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 122 | have "Pi\<^sub>E I (Y U) = U" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 123 | using Y \<open>U \<in> \<U>\<close> by blast | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 124 | then show "f \<in> U" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 125 | using that unfolding PiE_def Pi_def by blast | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 126 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 127 | show "?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U) \<subseteq> \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 128 | by (auto simp: PiE_iff *) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 129 | show "\<Inter>\<U> \<subseteq> ?TOP \<inter> \<Inter>(\<Union>U\<in>\<U>. ?F U)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 130 | using Y openin_subset \<open>finite' \<U>\<close> by fastforce | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 131 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 132 | qed (use Y openin_subset in \<open>blast+\<close>) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 133 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 134 | have 2: "\<exists>\<U>'. finite' \<U>' \<and> \<U>' \<subseteq> Collect ?\<Omega> \<and> \<Inter>\<U>' = ?TOP \<inter> \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 135 | if \<U>: "\<U> \<subseteq> Collect ?\<Psi>" and "finite \<U>" for \<U> | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 136 |     proof (cases "\<U>={}")
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 137 | case True | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 138 | then show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 139 |         apply (rule_tac x="{?TOP}" in exI, simp)
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 140 | apply (rule_tac x="\<lambda>i. topspace (X i)" in exI) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 141 | apply force | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 142 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 143 | next | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 144 | case False | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 145 | then obtain U where "U \<in> \<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 146 | by blast | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 147 |       have "\<forall>U \<in> \<U>. \<exists>i Y. U = {f. f i \<in> Y} \<and> i \<in> I \<and> openin (X i) Y"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 148 | using \<U> by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 149 | then obtain J Y where | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 150 |         Y: "\<And>U. U \<in> \<U> \<Longrightarrow> U = {f. f (J U) \<in> Y U} \<and> J U \<in> I \<and> openin (X (J U)) (Y U)"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 151 | by metis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 152 | let ?G = "\<lambda>U. \<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 153 | show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 154 | proof (intro conjI exI) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 155 |         show "finite (?G ` \<U>)" "?G ` \<U> \<noteq> {}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 156 | using \<open>finite \<U>\<close> \<open>U \<in> \<U>\<close> by blast+ | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 157 | have *: "\<And>U. U \<in> \<U> \<Longrightarrow> openin (X (J U)) (Y U)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 158 | using Y by force | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 159 |         show "?G ` \<U> \<subseteq> {Pi\<^sub>E I Y |Y. (\<forall>i. openin (X i) (Y i)) \<and> finite {i. Y i \<noteq> topspace (X i)}}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 160 | apply clarsimp | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 161 | apply (rule_tac x=" (\<lambda>i. if i = J U then Y U else topspace (X i))" in exI) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 162 | apply (auto simp: *) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 163 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 164 | next | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 165 | show "(\<Inter>U\<in>\<U>. ?G U) = ?TOP \<inter> \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 166 | proof | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 167 | have "(\<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 168 | apply (clarsimp simp: PiE_iff Ball_def all_conj_distrib split: if_split_asm) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 169 | using Y \<open>U \<in> \<U>\<close> openin_subset by blast+ | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 170 | then have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 171 | using \<open>U \<in> \<U>\<close> | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 172 | by fastforce | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 173 | moreover have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 174 | using PiE_mem Y by fastforce | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 175 | ultimately show "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP \<inter> \<Inter>\<U>" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 176 | by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 177 | qed (use Y in fastforce) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 178 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 179 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 180 | show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 181 | unfolding relative_to_def intersection_of_def | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 182 | by (safe; blast dest!: 1 2) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 183 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 184 | show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 185 | unfolding product_topology_def generate_topology_on_eq | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 186 | apply (rule arg_cong [where f = topology]) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 187 | apply (rule arg_cong [where f = "(union_of)arbitrary"]) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 188 | apply (force simp: *) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 189 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 190 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 191 | |
| 69874 | 192 | lemma topspace_product_topology [simp]: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 193 | "topspace (product_topology T I) = (\<Pi>\<^sub>E i\<in>I. topspace(T i))" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 194 | proof | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 195 | show "topspace (product_topology T I) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (T i))" | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 196 | unfolding product_topology_def topology_generated_by_topspace | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 197 | unfolding topspace_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 198 |   have "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<in> {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 199 | using openin_topspace not_finite_existsD by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 200 | then show "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<subseteq> topspace (product_topology T I)" | 
| 71633 | 201 | unfolding product_topology_def using PiE_def by (auto) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 202 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 203 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 204 | lemma product_topology_basis: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 205 |   assumes "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 206 | shows "openin (product_topology T I) (\<Pi>\<^sub>E i\<in>I. X i)" | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 207 | unfolding product_topology_def | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 208 | by (rule topology_generated_by_Basis) (use assms in auto) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 209 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 210 | proposition product_topology_open_contains_basis: | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 211 | assumes "openin (product_topology T I) U" "x \<in> U" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 212 |   shows "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 213 | proof - | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 214 |   have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}} U"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 215 | using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 216 |   then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 217 | proof induction | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 218 | case (Int U V x) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 219 | then obtain XU XV where H: | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 220 |       "x \<in> Pi\<^sub>E I XU" "(\<forall>i. openin (T i) (XU i))" "finite {i. XU i \<noteq> topspace (T i)}" "Pi\<^sub>E I XU \<subseteq> U"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 221 |       "x \<in> Pi\<^sub>E I XV" "(\<forall>i. openin (T i) (XV i))" "finite {i. XV i \<noteq> topspace (T i)}" "Pi\<^sub>E I XV \<subseteq> V"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 222 | by auto meson | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 223 | define X where "X = (\<lambda>i. XU i \<inter> XV i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 224 | have "Pi\<^sub>E I X \<subseteq> Pi\<^sub>E I XU \<inter> Pi\<^sub>E I XV" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 225 | unfolding X_def by (auto simp add: PiE_iff) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 226 | then have "Pi\<^sub>E I X \<subseteq> U \<inter> V" using H by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 227 | moreover have "\<forall>i. openin (T i) (X i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 228 | unfolding X_def using H by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 229 |     moreover have "finite {i. X i \<noteq> topspace (T i)}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 230 |       apply (rule rev_finite_subset[of "{i. XU i \<noteq> topspace (T i)} \<union> {i. XV i \<noteq> topspace (T i)}"])
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 231 | unfolding X_def using H by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 232 | moreover have "x \<in> Pi\<^sub>E I X" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 233 | unfolding X_def using H by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 234 | ultimately show ?case | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 235 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 236 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 237 | case (UN K x) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 238 | then obtain k where "k \<in> K" "x \<in> k" by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 239 |     with UN have "\<exists>X. x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 240 | by simp | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 241 |     then obtain X where "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 242 | by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 243 |     then have "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> (\<Union>K)"
 | 
| 64911 | 244 | using \<open>k \<in> K\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 245 | then show ?case | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 246 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 247 | qed auto | 
| 64911 | 248 | then show ?thesis using \<open>x \<in> U\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 249 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 250 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 251 | lemma product_topology_empty_discrete: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 252 |    "product_topology T {} = discrete_topology {(\<lambda>x. undefined)}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 253 | by (simp add: subtopology_eq_discrete_topology_sing) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 254 | |
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 255 | lemma openin_product_topology: | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 256 | "openin (product_topology X I) = | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 257 | arbitrary union_of | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 258 |           ((finite intersection_of (\<lambda>F. (\<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)))
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 259 | relative_to topspace (product_topology X I))" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 260 | apply (subst product_topology) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 261 | apply (simp add: topology_inverse' [OF istopology_subbase]) | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 262 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 263 | |
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 264 | lemma subtopology_PiE: | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 265 | "subtopology (product_topology X I) (\<Pi>\<^sub>E i\<in>I. (S i)) = product_topology (\<lambda>i. subtopology (X i) (S i)) I" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 266 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 267 |   let ?P = "\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 268 | let ?X = "\<Pi>\<^sub>E i\<in>I. topspace (X i)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 269 | have "finite intersection_of ?P relative_to ?X \<inter> Pi\<^sub>E I S = | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 270 | finite intersection_of (?P relative_to ?X \<inter> Pi\<^sub>E I S) relative_to ?X \<inter> Pi\<^sub>E I S" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 271 | by (rule finite_intersection_of_relative_to) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 272 | also have "\<dots> = finite intersection_of | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 273 |                       ((\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> (openin (X i) relative_to S i) U)
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 274 | relative_to ?X \<inter> Pi\<^sub>E I S) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 275 | relative_to ?X \<inter> Pi\<^sub>E I S" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 276 | apply (rule arg_cong2 [where f = "(relative_to)"]) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 277 | apply (rule arg_cong [where f = "(intersection_of)finite"]) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 278 | apply (rule ext) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 279 | apply (auto simp: relative_to_def intersection_of_def) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 280 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 281 | finally | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 282 | have "finite intersection_of ?P relative_to ?X \<inter> Pi\<^sub>E I S = | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 283 | finite intersection_of | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 284 |           (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> (openin (X i) relative_to S i) U)
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 285 | relative_to ?X \<inter> Pi\<^sub>E I S" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 286 | by (metis finite_intersection_of_relative_to) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 287 | then show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 288 | unfolding topology_eq | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 289 | apply clarify | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 290 | apply (simp add: openin_product_topology flip: openin_relative_to) | 
| 71172 | 291 | apply (simp add: arbitrary_union_of_relative_to flip: PiE_Int) | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 292 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 293 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 294 | |
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 295 | lemma product_topology_base_alt: | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 296 |    "finite intersection_of (\<lambda>F. (\<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U))
 | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 297 | relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 298 |     (\<lambda>F. (\<exists>U. F =  Pi\<^sub>E I U \<and> finite {i \<in> I. U i \<noteq> topspace(X i)} \<and> (\<forall>i \<in> I. openin (X i) (U i))))"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 299 | (is "?lhs = ?rhs") | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 300 | proof - | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 301 | have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 302 | unfolding all_relative_to all_intersection_of topspace_product_topology | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 303 | proof clarify | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 304 | fix \<F> | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 305 |     assume "finite \<F>" and "\<F> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (X i) U}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 306 | then show "\<exists>U. (\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>\<F> = Pi\<^sub>E I U \<and> | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 307 |           finite {i \<in> I. U i \<noteq> topspace (X i)} \<and> (\<forall>i\<in>I. openin (X i) (U i))"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 308 | proof (induction) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 309 | case (insert F \<F>) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 310 | then obtain U where eq: "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>\<F> = Pi\<^sub>E I U" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 311 |         and fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 312 | and ope: "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 313 | by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 314 |       obtain i V where "F = {f. f i \<in> V}" "i \<in> I" "openin (X i) V"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 315 | using insert by auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 316 | let ?U = "\<lambda>j. U j \<inter> (if j = i then V else topspace(X j))" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 317 | show ?case | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 318 | proof (intro exI conjI) | 
| 69745 | 319 | show "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> \<Inter>(insert F \<F>) = Pi\<^sub>E I ?U" | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 320 |         using eq  PiE_mem \<open>i \<in> I\<close>  by (auto simp: \<open>F = {f. f i \<in> V}\<close>) fastforce
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 321 | next | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 322 |         show "finite {i \<in> I. ?U i \<noteq> topspace (X i)}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 323 | by (rule rev_finite_subset [OF finite.insertI [OF fin]]) auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 324 | next | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 325 | show "\<forall>i\<in>I. openin (X i) (?U i)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 326 | by (simp add: \<open>openin (X i) V\<close> ope openin_Int) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 327 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 328 | qed (auto intro: dest: not_finite_existsD) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 329 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 330 | moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 331 | proof clarify | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 332 | fix U :: "'a \<Rightarrow> 'b set" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 333 |     assume fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}" and ope: "\<forall>i\<in>I. openin (X i) (U i)"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 334 |     let ?U = "\<Inter>i\<in>{i \<in> I. U i \<noteq> topspace (X i)}. {x. x i \<in> U i}"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 335 | show "?lhs (Pi\<^sub>E I U)" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 336 | unfolding relative_to_def topspace_product_topology | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 337 | proof (intro exI conjI) | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 338 |       show "(finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)) ?U"
 | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 339 | using fin ope by (intro finite_intersection_of_Inter finite_intersection_of_inc) auto | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 340 | show "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> ?U = Pi\<^sub>E I U" | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 341 | using ope openin_subset by fastforce | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 342 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 343 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 344 | ultimately show ?thesis | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 345 | by meson | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 346 | qed | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 347 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 348 | corollary openin_product_topology_alt: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 349 | "openin (product_topology X I) S \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 350 |     (\<forall>x \<in> S. \<exists>U. finite {i \<in> I. U i \<noteq> topspace(X i)} \<and>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 351 | (\<forall>i \<in> I. openin (X i) (U i)) \<and> x \<in> Pi\<^sub>E I U \<and> Pi\<^sub>E I U \<subseteq> S)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 352 | unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 353 | apply safe | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 354 | apply (drule bspec; blast)+ | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 355 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 356 | |
| 69874 | 357 | lemma closure_of_product_topology: | 
| 358 | "(product_topology X I) closure_of (PiE I S) = PiE I (\<lambda>i. (X i) closure_of (S i))" | |
| 359 | proof - | |
| 360 | have *: "(\<forall>T. f \<in> T \<and> openin (product_topology X I) T \<longrightarrow> (\<exists>y\<in>Pi\<^sub>E I S. y \<in> T)) | |
| 361 |        \<longleftrightarrow> (\<forall>i \<in> I. \<forall>T. f i \<in> T \<and> openin (X i) T \<longrightarrow> S i \<inter> T \<noteq> {})"
 | |
| 362 | (is "?lhs = ?rhs") | |
| 363 | if top: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i)" and ext: "f \<in> extensional I" for f | |
| 364 | proof | |
| 365 | assume L[rule_format]: ?lhs | |
| 366 | show ?rhs | |
| 367 | proof clarify | |
| 368 | fix i T | |
| 369 |       assume "i \<in> I" "f i \<in> T" "openin (X i) T" "S i \<inter> T = {}"
 | |
| 370 |       then have "openin (product_topology X I) ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> {x. x i \<in> T})"
 | |
| 371 | by (force simp: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc) | |
| 372 | then show "False" | |
| 373 |         using L [of "topspace (product_topology X I) \<inter> {f. f i \<in> T}"] \<open>S i \<inter> T = {}\<close> \<open>f i \<in> T\<close> \<open>i \<in> I\<close>
 | |
| 374 | by (auto simp: top ext PiE_iff) | |
| 375 | qed | |
| 376 | next | |
| 377 | assume R [rule_format]: ?rhs | |
| 378 | show ?lhs | |
| 379 | proof (clarsimp simp: openin_product_topology union_of_def arbitrary_def) | |
| 380 | fix \<U> U | |
| 381 | assume | |
| 382 | \<U>: "\<U> \<subseteq> Collect | |
| 383 |            (finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U) relative_to
 | |
| 384 | (\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and | |
| 385 | "f \<in> U" "U \<in> \<U>" | |
| 386 |       then have "(finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U)
 | |
| 387 | relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))) U" | |
| 388 | by blast | |
| 389 | with \<open>f \<in> U\<close> \<open>U \<in> \<U>\<close> | |
| 390 | obtain \<T> where "finite \<T>" | |
| 391 |         and \<T>: "\<And>C. C \<in> \<T> \<Longrightarrow> \<exists>i \<in> I. \<exists>V. openin (X i) V \<and> C = {x. x i \<in> V}"
 | |
| 392 | and "topspace (product_topology X I) \<inter> \<Inter> \<T> \<subseteq> U" "f \<in> topspace (product_topology X I) \<inter> \<Inter> \<T>" | |
| 393 | apply (clarsimp simp add: relative_to_def intersection_of_def) | |
| 394 | apply (rule that, auto dest!: subsetD) | |
| 395 | done | |
| 396 | then have "f \<in> PiE I (topspace \<circ> X)" "f \<in> \<Inter>\<T>" and subU: "PiE I (topspace \<circ> X) \<inter> \<Inter>\<T> \<subseteq> U" | |
| 397 | by (auto simp: PiE_iff) | |
| 398 |       have *: "f i \<in> topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}
 | |
| 399 |              \<and> openin (X i) (topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>})"
 | |
| 400 | if "i \<in> I" for i | |
| 401 | proof - | |
| 402 |         have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)"
 | |
| 403 | proof (rule finite_vimageI [OF \<open>finite \<T>\<close>]) | |
| 404 |           show "inj (\<lambda>U. {x. x i \<in> U})"
 | |
| 405 | by (auto simp: inj_on_def) | |
| 406 | qed | |
| 407 |         then have fin: "finite {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}"
 | |
| 408 | by (rule rev_finite_subset) auto | |
| 409 |         have "openin (X i) (\<Inter> (insert (topspace (X i)) {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}))"
 | |
| 410 | by (rule openin_Inter) (auto simp: fin) | |
| 411 | then show ?thesis | |
| 412 | using \<open>f \<in> \<Inter> \<T>\<close> by (fastforce simp: that top) | |
| 413 | qed | |
| 414 |       define \<Phi> where "\<Phi> \<equiv> \<lambda>i. topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {f. f i \<in> U} \<in> \<T>}"
 | |
| 415 | have "\<forall>i \<in> I. \<exists>x. x \<in> S i \<inter> \<Phi> i" | |
| 416 | using R [OF _ *] unfolding \<Phi>_def by blast | |
| 417 | then obtain \<theta> where \<theta> [rule_format]: "\<forall>i \<in> I. \<theta> i \<in> S i \<inter> \<Phi> i" | |
| 418 | by metis | |
| 419 | show "\<exists>y\<in>Pi\<^sub>E I S. \<exists>x\<in>\<U>. y \<in> x" | |
| 420 | proof | |
| 421 | show "\<exists>U \<in> \<U>. (\<lambda>i \<in> I. \<theta> i) \<in> U" | |
| 422 | proof | |
| 423 | have "restrict \<theta> I \<in> PiE I (topspace \<circ> X) \<inter> \<Inter>\<T>" | |
| 424 | using \<T> by (fastforce simp: \<Phi>_def PiE_def dest: \<theta>) | |
| 425 | then show "restrict \<theta> I \<in> U" | |
| 426 | using subU by blast | |
| 427 | qed (rule \<open>U \<in> \<U>\<close>) | |
| 428 | next | |
| 429 | show "(\<lambda>i \<in> I. \<theta> i) \<in> Pi\<^sub>E I S" | |
| 430 | using \<theta> by simp | |
| 431 | qed | |
| 432 | qed | |
| 433 | qed | |
| 434 | show ?thesis | |
| 435 | apply (simp add: * closure_of_def PiE_iff set_eq_iff cong: conj_cong) | |
| 436 | by metis | |
| 437 | qed | |
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 438 | |
| 69874 | 439 | corollary closedin_product_topology: | 
| 440 |    "closedin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. closedin (X i) (S i))"
 | |
| 441 | apply (simp add: PiE_eq PiE_eq_empty_iff closure_of_product_topology flip: closure_of_eq) | |
| 442 | apply (metis closure_of_empty) | |
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 443 | done | 
| 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 444 | |
| 69874 | 445 | corollary closedin_product_topology_singleton: | 
| 446 |    "f \<in> extensional I \<Longrightarrow> closedin (product_topology X I) {f} \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) {f i})"
 | |
| 447 | using PiE_singleton closedin_product_topology [of X I] | |
| 448 | by (metis (no_types, lifting) all_not_in_conv insertI1) | |
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 449 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 450 | lemma product_topology_empty: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 451 |    "product_topology X {} = topology (\<lambda>S. S \<in> {{},{\<lambda>k. undefined}})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 452 | unfolding product_topology union_of_def intersection_of_def arbitrary_def relative_to_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 453 | by (auto intro: arg_cong [where f=topology]) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 454 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 455 | lemma openin_product_topology_empty: "openin (product_topology X {}) S \<longleftrightarrow> S \<in> {{},{\<lambda>k. undefined}}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 456 | unfolding union_of_def intersection_of_def arbitrary_def relative_to_def openin_product_topology | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 457 | by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 458 | |
| 69874 | 459 | |
| 460 | subsubsection \<open>The basic property of the product topology is the continuity of projections:\<close> | |
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 461 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 462 | lemma continuous_map_product_coordinates [simp]: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 463 | assumes "i \<in> I" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 464 | shows "continuous_map (product_topology T I) (T i) (\<lambda>x. x i)" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 465 | proof - | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 466 |   {
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 467 | fix U assume "openin (T i) U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 468 | define X where "X = (\<lambda>j. if j = i then U else topspace (T j))" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 469 | then have *: "(\<lambda>x. x i) -` U \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (T i)) = (\<Pi>\<^sub>E j\<in>I. X j)" | 
| 64911 | 470 | unfolding X_def using assms openin_subset[OF \<open>openin (T i) U\<close>] | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 471 | by (auto simp add: PiE_iff, auto, metis subsetCE) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 472 |     have **: "(\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}"
 | 
| 64911 | 473 | unfolding X_def using \<open>openin (T i) U\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 474 | have "openin (product_topology T I) ((\<lambda>x. x i) -` U \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (T i)))" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 475 | unfolding product_topology_def | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 476 | apply (rule topology_generated_by_Basis) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 477 | apply (subst *) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 478 | using ** by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 479 | } | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 480 | then show ?thesis unfolding continuous_map_alt | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 481 | by (auto simp add: assms PiE_iff) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 482 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 483 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 484 | lemma continuous_map_coordinatewise_then_product [intro]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 485 | assumes "\<And>i. i \<in> I \<Longrightarrow> continuous_map T1 (T i) (\<lambda>x. f x i)" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 486 | "\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 487 | shows "continuous_map T1 (product_topology T I) f" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 488 | unfolding product_topology_def | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 489 | proof (rule continuous_on_generated_topo) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 490 |   fix U assume "U \<in> {Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 491 |   then obtain X where H: "U = Pi\<^sub>E I X" "\<And>i. openin (T i) (X i)" "finite {i. X i \<noteq> topspace (T i)}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 492 | by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 493 |   define J where "J = {i \<in> I. X i \<noteq> topspace (T i)}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 494 | have "finite J" "J \<subseteq> I" unfolding J_def using H(3) by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 495 | have "(\<lambda>x. f x i)-`(topspace(T i)) \<inter> topspace T1 = topspace T1" if "i \<in> I" for i | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 496 | using that assms(1) by (simp add: continuous_map_preimage_topspace) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 497 | then have *: "(\<lambda>x. f x i)-`(X i) \<inter> topspace T1 = topspace T1" if "i \<in> I-J" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 498 | using that unfolding J_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 499 | have "f-`U \<inter> topspace T1 = (\<Inter>i\<in>I. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 500 | by (subst H(1), auto simp add: PiE_iff assms) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 501 | also have "... = (\<Inter>i\<in>J. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)" | 
| 64911 | 502 | using * \<open>J \<subseteq> I\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 503 | also have "openin T1 (...)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 504 | apply (rule openin_INT) | 
| 64911 | 505 | apply (simp add: \<open>finite J\<close>) | 
| 506 | using H(2) assms(1) \<open>J \<subseteq> I\<close> by auto | |
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 507 | ultimately show "openin T1 (f-`U \<inter> topspace T1)" by simp | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 508 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 509 |   show "f `topspace T1 \<subseteq> \<Union>{Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 510 | apply (subst topology_generated_by_topspace[symmetric]) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 511 | apply (subst product_topology_def[symmetric]) | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 512 | apply (simp only: topspace_product_topology) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 513 | apply (auto simp add: PiE_iff) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 514 | using assms unfolding continuous_map_def by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 515 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 516 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 517 | lemma continuous_map_product_then_coordinatewise [intro]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 518 | assumes "continuous_map T1 (product_topology T I) f" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 519 | shows "\<And>i. i \<in> I \<Longrightarrow> continuous_map T1 (T i) (\<lambda>x. f x i)" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 520 | "\<And>i x. i \<notin> I \<Longrightarrow> x \<in> topspace T1 \<Longrightarrow> f x i = undefined" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 521 | proof - | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 522 | fix i assume "i \<in> I" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 523 | have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f" by auto | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 524 | also have "continuous_map T1 (T i) (...)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 525 | apply (rule continuous_map_compose[of _ "product_topology T I"]) | 
| 64911 | 526 | using assms \<open>i \<in> I\<close> by auto | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 527 | ultimately show "continuous_map T1 (T i) (\<lambda>x. f x i)" | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 528 | by simp | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 529 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 530 | fix i x assume "i \<notin> I" "x \<in> topspace T1" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 531 | have "f x \<in> topspace (product_topology T I)" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 532 | using assms \<open>x \<in> topspace T1\<close> unfolding continuous_map_def by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 533 | then have "f x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (T i))" | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 534 | using topspace_product_topology by metis | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 535 | then show "f x i = undefined" | 
| 64911 | 536 | using \<open>i \<notin> I\<close> by (auto simp add: PiE_iff extensional_def) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 537 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 538 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 539 | lemma continuous_on_restrict: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 540 | assumes "J \<subseteq> I" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 541 | shows "continuous_map (product_topology T I) (product_topology T J) (\<lambda>x. restrict x J)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 542 | proof (rule continuous_map_coordinatewise_then_product) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 543 | fix i assume "i \<in> J" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 544 | then have "(\<lambda>x. restrict x J i) = (\<lambda>x. x i)" unfolding restrict_def by auto | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 545 | then show "continuous_map (product_topology T I) (T i) (\<lambda>x. restrict x J i)" | 
| 64911 | 546 | using \<open>i \<in> J\<close> \<open>J \<subseteq> I\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 547 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 548 | fix i assume "i \<notin> J" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 549 | then show "restrict x J i = undefined" for x::"'a \<Rightarrow> 'b" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 550 | unfolding restrict_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 551 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 552 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 553 | |
| 70136 | 554 | subsubsection\<^marker>\<open>tag important\<close> \<open>Powers of a single topological space as a topological space, using type classes\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 555 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 556 | instantiation "fun" :: (type, topological_space) topological_space | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 557 | begin | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 558 | |
| 70136 | 559 | definition\<^marker>\<open>tag important\<close> open_fun_def: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 560 | "open U = openin (product_topology (\<lambda>i. euclidean) UNIV) U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 561 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 562 | instance proof | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 563 |   have "topspace (product_topology (\<lambda>(i::'a). euclidean::('b topology)) UNIV) = UNIV"
 | 
| 69030 
1eb517214deb
more on product (function) topologies
 paulson <lp15@cam.ac.uk> parents: 
68833diff
changeset | 564 | unfolding topspace_product_topology topspace_euclidean by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 565 |   then show "open (UNIV::('a \<Rightarrow> 'b) set)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 566 | unfolding open_fun_def by (metis openin_topspace) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 567 | qed (auto simp add: open_fun_def) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 568 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 569 | end | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 570 | |
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 571 | lemma open_PiE [intro?]: | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 572 |   fixes X::"'i \<Rightarrow> ('b::topological_space) set"
 | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 573 |   assumes "\<And>i. open (X i)" "finite {i. X i \<noteq> UNIV}"
 | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 574 | shows "open (Pi\<^sub>E UNIV X)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 575 | by (simp add: assms open_fun_def product_topology_basis) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 576 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 577 | lemma euclidean_product_topology: | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 578 |   "product_topology (\<lambda>i. euclidean::('b::topological_space) topology) UNIV = euclidean"
 | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 579 | by (metis open_openin topology_eq open_fun_def) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 580 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 581 | proposition product_topology_basis': | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 582 |   fixes x::"'i \<Rightarrow> 'a" and U::"'i \<Rightarrow> ('b::topological_space) set"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 583 | assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> open (U i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 584 |   shows "open {f. \<forall>i\<in>I. f (x i) \<in> U i}"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 585 | proof - | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 586 | define J where "J = x`I" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 587 |   define V where "V = (\<lambda>y. if y \<in> J then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 588 | define X where "X = (\<lambda>y. if y \<in> J then V y else UNIV)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 589 | have *: "open (X i)" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 590 | unfolding X_def V_def using assms by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 591 |   have **: "finite {i. X i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 592 | unfolding X_def V_def J_def using assms(1) by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 593 | have "open (Pi\<^sub>E UNIV X)" | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 594 | by (simp add: "*" "**" open_PiE) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 595 |   moreover have "Pi\<^sub>E UNIV X = {f. \<forall>i\<in>I. f (x i) \<in> U i}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 596 | apply (auto simp add: PiE_iff) unfolding X_def V_def J_def | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 597 | proof (auto) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 598 | fix f :: "'a \<Rightarrow> 'b" and i :: 'i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 599 | assume a1: "i \<in> I" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 600 |       assume a2: "\<forall>i. f i \<in> (if i \<in> x`I then if i \<in> x`I then \<Inter>{U ia |ia. ia \<in> I \<and> x ia = i} else UNIV else UNIV)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 601 | have f3: "x i \<in> x`I" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 602 | using a1 by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 603 |       have "U i \<in> {U ia |ia. ia \<in> I \<and> x ia = x i}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 604 | using a1 by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 605 | then show "f (x i) \<in> U i" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 606 | using f3 a2 by (meson Inter_iff) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 607 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 608 | ultimately show ?thesis by simp | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 609 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 610 | |
| 64911 | 611 | text \<open>The results proved in the general situation of products of possibly different | 
| 612 | spaces have their counterparts in this simpler setting.\<close> | |
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 613 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 614 | lemma continuous_on_product_coordinates [simp]: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 615 |   "continuous_on UNIV (\<lambda>x. x i::('b::topological_space))"
 | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 616 | using continuous_map_product_coordinates [of _ UNIV "\<lambda>i. euclidean"] | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 617 | by (metis (no_types) continuous_map_iff_continuous euclidean_product_topology iso_tuple_UNIV_I subtopology_UNIV) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 618 | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 619 | lemma continuous_on_coordinatewise_then_product [continuous_intros]: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 620 | fixes f :: "'a::topological_space \<Rightarrow> 'b \<Rightarrow> 'c::topological_space" | 
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 621 | assumes "\<And>i. continuous_on S (\<lambda>x. f x i)" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 622 | shows "continuous_on S f" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 623 | using continuous_map_coordinatewise_then_product [of UNIV, where T = "\<lambda>i. euclidean"] | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 624 | by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 625 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 626 | lemma continuous_on_product_then_coordinatewise_UNIV: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 627 | assumes "continuous_on UNIV f" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 628 | shows "continuous_on UNIV (\<lambda>x. f x i)" | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 629 | unfolding continuous_map_iff_continuous2 [symmetric] | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 630 | by (rule continuous_map_product_then_coordinatewise [where I=UNIV]) (use assms in \<open>auto simp: euclidean_product_topology\<close>) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 631 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 632 | lemma continuous_on_product_then_coordinatewise: | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 633 | assumes "continuous_on S f" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 634 | shows "continuous_on S (\<lambda>x. f x i)" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 635 | proof - | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 636 | have "continuous_on S ((\<lambda>q. q i) \<circ> f)" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 637 | by (metis assms continuous_on_compose continuous_on_id | 
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 638 | continuous_on_product_then_coordinatewise_UNIV continuous_on_subset subset_UNIV) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 639 | then show ?thesis | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 640 | by auto | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 641 | qed | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 642 | |
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 643 | lemma continuous_on_coordinatewise_iff: | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 644 |   fixes f :: "('a \<Rightarrow> real) \<Rightarrow> 'b \<Rightarrow> real"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 645 | shows "continuous_on (A \<inter> S) f \<longleftrightarrow> (\<forall>i. continuous_on (A \<inter> S) (\<lambda>x. f x i))" | 
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 646 | by (auto simp: continuous_on_product_then_coordinatewise continuous_on_coordinatewise_then_product) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 647 | |
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 648 | lemma continuous_map_span_sum: | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71172diff
changeset | 649 | fixes B :: "'a::real_normed_vector set" | 
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 650 | assumes biB: "\<And>i. i \<in> I \<Longrightarrow> b i \<in> B" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 651 | shows "continuous_map euclidean (top_of_set (span B)) (\<lambda>x. \<Sum>i\<in>I. x i *\<^sub>R b i)" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 652 | proof (rule continuous_map_euclidean_top_of_set) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 653 | show "(\<lambda>x. \<Sum>i\<in>I. x i *\<^sub>R b i) -` span B = UNIV" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 654 | by auto (meson biB lessThan_iff span_base span_scale span_sum) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 655 | show "continuous_on UNIV (\<lambda>x. \<Sum>i\<in> I. x i *\<^sub>R b i)" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 656 | by (intro continuous_intros) auto | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 657 | qed | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 658 | |
| 70136 | 659 | subsubsection\<^marker>\<open>tag important\<close> \<open>Topological countability for product spaces\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 660 | |
| 64911 | 661 | text \<open>The next two lemmas are useful to prove first or second countability | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 662 | of product spaces, but they have more to do with countability and could | 
| 64911 | 663 | be put in the corresponding theory.\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 664 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 665 | lemma countable_nat_product_event_const: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 666 | fixes F::"'a set" and a::'a | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 667 | assumes "a \<in> F" "countable F" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 668 |   shows "countable {x::(nat \<Rightarrow> 'a). (\<forall>i. x i \<in> F) \<and> finite {i. x i \<noteq> a}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 669 | proof - | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 670 |   have *: "{x::(nat \<Rightarrow> 'a). (\<forall>i. x i \<in> F) \<and> finite {i. x i \<noteq> a}}
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 671 |                   \<subseteq> (\<Union>N. {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)})"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 672 | using infinite_nat_iff_unbounded_le by fastforce | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 673 |   have "countable {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)}" for N::nat
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 674 | proof (induction N) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 675 | case 0 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 676 |     have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>(0::nat). x i = a)} = {(\<lambda>i. a)}"
 | 
| 64911 | 677 | using \<open>a \<in> F\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 678 | then show ?case by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 679 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 680 | case (Suc N) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 681 | define f::"((nat \<Rightarrow> 'a) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 682 | where "f = (\<lambda>(x, b). (\<lambda>i. if i = N then b else x i))" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 683 |     have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>Suc N. x i = a)} \<subseteq> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 684 | proof (auto) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 685 | fix x assume H: "\<forall>i::nat. x i \<in> F" "\<forall>i\<ge>Suc N. x i = a" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 686 | define y where "y = (\<lambda>i. if i = N then a else x i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 687 | have "f (y, x N) = x" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 688 | unfolding f_def y_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 689 |       moreover have "(y, x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F"
 | 
| 64911 | 690 | unfolding y_def using H \<open>a \<in> F\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 691 |       ultimately show "x \<in> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 692 | by (metis (no_types, lifting) image_eqI) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 693 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 694 |     moreover have "countable ({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 695 | using Suc.IH assms(2) by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 696 | ultimately show ?case | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 697 | by (meson countable_image countable_subset) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 698 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 699 | then show ?thesis using countable_subset[OF *] by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 700 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 701 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 702 | lemma countable_product_event_const: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 703 |   fixes F::"('a::countable) \<Rightarrow> 'b set" and b::'b
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 704 | assumes "\<And>i. countable (F i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 705 |   shows "countable {f::('a \<Rightarrow> 'b). (\<forall>i. f i \<in> F i) \<and> (finite {i. f i \<noteq> b})}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 706 | proof - | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 707 |   define G where "G = (\<Union>i. F i) \<union> {b}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 708 | have "countable G" unfolding G_def using assms by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 709 | have "b \<in> G" unfolding G_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 710 |   define pi where "pi = (\<lambda>(x::(nat \<Rightarrow> 'b)). (\<lambda> i::'a. x ((to_nat::('a \<Rightarrow> nat)) i)))"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 711 |   have "{f::('a \<Rightarrow> 'b). (\<forall>i. f i \<in> F i) \<and> (finite {i. f i \<noteq> b})}
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 712 |         \<subseteq> pi`{g::(nat \<Rightarrow> 'b). (\<forall>j. g j \<in> G) \<and> (finite {j. g j \<noteq> b})}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 713 | proof (auto) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 714 |     fix f assume H: "\<forall>i. f i \<in> F i" "finite {i. f i \<noteq> b}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 715 |     define I where "I = {i. f i \<noteq> b}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 716 | define g where "g = (\<lambda>j. if j \<in> to_nat`I then f (from_nat j) else b)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 717 |     have "{j. g j \<noteq> b} \<subseteq> to_nat`I" unfolding g_def by auto
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 718 |     then have "finite {j. g j \<noteq> b}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 719 | unfolding I_def using H(2) using finite_surj by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 720 | moreover have "g j \<in> G" for j | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 721 | unfolding g_def G_def using H by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 722 |     ultimately have "g \<in> {g::(nat \<Rightarrow> 'b). (\<forall>j. g j \<in> G) \<and> (finite {j. g j \<noteq> b})}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 723 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 724 | moreover have "f = pi g" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 725 | unfolding pi_def g_def I_def using H by fastforce | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 726 |     ultimately show "f \<in> pi`{g. (\<forall>j. g j \<in> G) \<and> finite {j. g j \<noteq> b}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 727 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 728 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 729 | then show ?thesis | 
| 64911 | 730 | using countable_nat_product_event_const[OF \<open>b \<in> G\<close> \<open>countable G\<close>] | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 731 | by (meson countable_image countable_subset) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 732 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 733 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 734 | instance "fun" :: (countable, first_countable_topology) first_countable_topology | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 735 | proof | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 736 | fix x::"'a \<Rightarrow> 'b" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 737 |   have "\<exists>A::('b \<Rightarrow> nat \<Rightarrow> 'b set). \<forall>x. (\<forall>i. x \<in> A x i \<and> open (A x i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A x i \<subseteq> S))"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 738 | apply (rule choice) using first_countable_basis by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 739 |   then obtain A::"('b \<Rightarrow> nat \<Rightarrow> 'b set)" where A: "\<And>x i. x \<in> A x i"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 740 | "\<And>x i. open (A x i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 741 | "\<And>x S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>i. A x i \<subseteq> S)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 742 | by metis | 
| 69566 | 743 | text \<open>\<open>B i\<close> is a countable basis of neighborhoods of \<open>x\<^sub>i\<close>.\<close> | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 744 |   define B where "B = (\<lambda>i. (A (x i))`UNIV \<union> {UNIV})"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 745 | have "countable (B i)" for i unfolding B_def by auto | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 746 | have open_B: "\<And>X i. X \<in> B i \<Longrightarrow> open X" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 747 | by (auto simp: B_def A) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 748 |   define K where "K = {Pi\<^sub>E UNIV X | X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 749 | have "Pi\<^sub>E UNIV (\<lambda>i. UNIV) \<in> K" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 750 | unfolding K_def B_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 751 |   then have "K \<noteq> {}" by auto
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 752 |   have "countable {X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 64911 | 753 | apply (rule countable_product_event_const) using \<open>\<And>i. countable (B i)\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 754 |   moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 755 | unfolding K_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 756 | ultimately have "countable K" by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 757 | have "x \<in> k" if "k \<in> K" for k | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 758 | using that unfolding K_def B_def apply auto using A(1) by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 759 | have "open k" if "k \<in> K" for k | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 760 | using that unfolding K_def by (blast intro: open_B open_PiE elim: ) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 761 | have Inc: "\<exists>k\<in>K. k \<subseteq> U" if "open U \<and> x \<in> U" for U | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 762 | proof - | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 763 | have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U" | 
| 64911 | 764 | using \<open>open U \<and> x \<in> U\<close> unfolding open_fun_def by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 765 | with product_topology_open_contains_basis[OF this] | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 766 |     have "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV} \<and> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
 | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69683diff
changeset | 767 | by simp | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 768 | then obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 769 | "\<And>i. open (X i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 770 |                            "finite {i. X i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 771 | "(\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 772 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 773 |     define I where "I = {i. X i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 774 | define Y where "Y = (\<lambda>i. if i \<in> I then (SOME y. y \<in> B i \<and> y \<subseteq> X i) else UNIV)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 775 | have *: "\<exists>y. y \<in> B i \<and> y \<subseteq> X i" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 776 | unfolding B_def using A(3)[OF H(2)] H(1) by (metis PiE_E UNIV_I UnCI image_iff) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 777 | have **: "Y i \<in> B i \<and> Y i \<subseteq> X i" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 778 | apply (cases "i \<in> I") | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 779 | unfolding Y_def using * that apply (auto) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 780 | apply (metis (no_types, lifting) someI, metis (no_types, lifting) someI_ex subset_iff) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 781 | unfolding B_def apply simp | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 782 | unfolding I_def apply auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 783 | done | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 784 |     have "{i. Y i \<noteq> UNIV} \<subseteq> I"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 785 | unfolding Y_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 786 |     then have ***: "finite {i. Y i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 787 | unfolding I_def using H(3) rev_finite_subset by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 788 |     have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 789 | using ** *** by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 790 | then have "Pi\<^sub>E UNIV Y \<in> K" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 791 | unfolding K_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 792 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 793 | have "Y i \<subseteq> X i" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 794 | apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 795 | then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 796 | then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto | 
| 64911 | 797 | then show ?thesis using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 798 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 799 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 800 | show "\<exists>L. (\<forall>(i::nat). x \<in> L i \<and> open (L i)) \<and> (\<forall>U. open U \<and> x \<in> U \<longrightarrow> (\<exists>i. L i \<subseteq> U))" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 801 | apply (rule first_countableI[of K]) | 
| 64911 | 802 | using \<open>countable K\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> x \<in> k\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> open k\<close> Inc by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 803 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 804 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 805 | proposition product_topology_countable_basis: | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 806 |   shows "\<exists>K::(('a::countable \<Rightarrow> 'b::second_countable_topology) set set).
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 807 | topological_basis K \<and> countable K \<and> | 
| 64910 | 808 |           (\<forall>k\<in>K. \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV})"
 | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 809 | proof - | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 810 | obtain B::"'b set set" where B: "countable B \<and> topological_basis B" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 811 | using ex_countable_basis by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 812 |   then have "B \<noteq> {}" by (meson UNIV_I empty_iff open_UNIV topological_basisE)
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 813 |   define B2 where "B2 = B \<union> {UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 814 | have "countable B2" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 815 | unfolding B2_def using B by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 816 | have "open U" if "U \<in> B2" for U | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 817 | using that unfolding B2_def using B topological_basis_open by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 818 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 819 |   define K where "K = {Pi\<^sub>E UNIV X | X. (\<forall>i::'a. X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 64910 | 820 |   have i: "\<forall>k\<in>K. \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV}"
 | 
| 64911 | 821 | unfolding K_def using \<open>\<And>U. U \<in> B2 \<Longrightarrow> open U\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 822 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 823 |   have "countable {X. (\<forall>(i::'a). X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 64911 | 824 | apply (rule countable_product_event_const) using \<open>countable B2\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 825 |   moreover have "K = (\<lambda>X. Pi\<^sub>E UNIV X)`{X. (\<forall>i. X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 826 | unfolding K_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 827 | ultimately have ii: "countable K" by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 828 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 829 | have iii: "topological_basis K" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 830 | proof (rule topological_basisI) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 831 | fix U and x::"'a\<Rightarrow>'b" assume "open U" "x \<in> U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 832 | then have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 833 | unfolding open_fun_def by auto | 
| 64911 | 834 | with product_topology_open_contains_basis[OF this \<open>x \<in> U\<close>] | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 835 |     have "\<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV} \<and> (\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U"
 | 
| 69710 
61372780515b
some renamings and a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
69683diff
changeset | 836 | by simp | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 837 | then obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. X i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 838 | "\<And>i. open (X i)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 839 |                            "finite {i. X i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 840 | "(\<Pi>\<^sub>E i\<in>UNIV. X i) \<subseteq> U" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 841 | by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 842 | then have "x i \<in> X i" for i by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 843 |     define I where "I = {i. X i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 844 | define Y where "Y = (\<lambda>i. if i \<in> I then (SOME y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y) else UNIV)" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 845 | have *: "\<exists>y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y" for i | 
| 64911 | 846 | unfolding B2_def using B \<open>open (X i)\<close> \<open>x i \<in> X i\<close> by (meson UnCI topological_basisE) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 847 | have **: "Y i \<in> B2 \<and> Y i \<subseteq> X i \<and> x i \<in> Y i" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 848 | using someI_ex[OF *] | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 849 | apply (cases "i \<in> I") | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 850 | unfolding Y_def using * apply (auto) | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 851 | unfolding B2_def I_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 852 |     have "{i. Y i \<noteq> UNIV} \<subseteq> I"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 853 | unfolding Y_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 854 |     then have ***: "finite {i. Y i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 855 | unfolding I_def using H(3) rev_finite_subset by blast | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 856 |     have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}"
 | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 857 | using ** *** by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 858 | then have "Pi\<^sub>E UNIV Y \<in> K" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 859 | unfolding K_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 860 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 861 | have "Y i \<subseteq> X i" for i | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 862 | apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 863 | then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 864 | then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 865 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 866 | have "x \<in> Pi\<^sub>E UNIV Y" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 867 | using ** by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 868 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 869 | show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U" | 
| 64911 | 870 | using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> \<open>Pi\<^sub>E UNIV Y \<subseteq> U\<close> \<open>x \<in> Pi\<^sub>E UNIV Y\<close> by auto | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 871 | next | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 872 | fix U assume "U \<in> K" | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 873 | show "open U" | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 874 | using \<open>U \<in> K\<close> unfolding open_fun_def K_def by clarify (metis \<open>U \<in> K\<close> i open_PiE open_fun_def) | 
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 875 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 876 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 877 | show ?thesis using i ii iii by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 878 | qed | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 879 | |
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 880 | instance "fun" :: (countable, second_countable_topology) second_countable_topology | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 881 | apply standard | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 882 | using product_topology_countable_basis topological_basis_imp_subbasis by auto | 
| 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 883 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 884 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 885 | subsection\<open>The Alexander subbase theorem\<close> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 886 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 887 | theorem Alexander_subbase: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 888 | assumes X: "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) = X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 889 | and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; \<Union>C = topspace X\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> \<Union>C' = topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 890 | shows "compact_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 891 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 892 | have U\<B>: "\<Union>\<B> = topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 893 | by (simp flip: X) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 894 | have False if \<U>: "\<forall>U\<in>\<U>. openin X U" and sub: "topspace X \<subseteq> \<Union>\<U>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 895 | and neg: "\<And>\<F>. \<lbrakk>\<F> \<subseteq> \<U>; finite \<F>\<rbrakk> \<Longrightarrow> \<not> topspace X \<subseteq> \<Union>\<F>" for \<U> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 896 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 897 |     define \<A> where "\<A> \<equiv> {\<C>. (\<forall>U \<in> \<C>. openin X U) \<and> topspace X \<subseteq> \<Union>\<C> \<and> (\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<C> \<longrightarrow> ~(topspace X \<subseteq> \<Union>\<F>))}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 898 |     have 1: "\<A> \<noteq> {}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 899 | unfolding \<A>_def using sub \<U> neg by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 900 |     have 2: "\<Union>\<C> \<in> \<A>" if "\<C>\<noteq>{}" and \<C>: "subset.chain \<A> \<C>" for \<C>
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 901 | unfolding \<A>_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 902 | proof (intro CollectI conjI ballI allI impI notI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 903 | show "openin X U" if U: "U \<in> \<Union>\<C>" for U | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 904 | using U \<C> unfolding \<A>_def subset_chain_def by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 905 | have "\<C> \<subseteq> \<A>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 906 | using subset_chain_def \<C> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 907 | with that \<A>_def show UUC: "topspace X \<subseteq> \<Union>(\<Union>\<C>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 908 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 909 | show "False" if "finite \<F>" and "\<F> \<subseteq> \<Union>\<C>" and "topspace X \<subseteq> \<Union>\<F>" for \<F> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 910 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 911 | obtain \<B> where "\<B> \<in> \<C>" "\<F> \<subseteq> \<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 912 | by (metis Sup_empty \<C> \<open>\<F> \<subseteq> \<Union>\<C>\<close> \<open>finite \<F>\<close> UUC empty_subsetI finite.emptyI finite_subset_Union_chain neg) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 913 | then show False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 914 | using \<A>_def \<open>\<C> \<subseteq> \<A>\<close> \<open>finite \<F>\<close> \<open>topspace X \<subseteq> \<Union>\<F>\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 915 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 916 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 917 | obtain \<K> where "\<K> \<in> \<A>" and "\<And>X. \<lbrakk>X\<in>\<A>; \<K> \<subseteq> X\<rbrakk> \<Longrightarrow> X = \<K>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 918 | using subset_Zorn_nonempty [OF 1 2] by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 919 | then have *: "\<And>\<W>. \<lbrakk>\<And>W. W\<in>\<W> \<Longrightarrow> openin X W; topspace X \<subseteq> \<Union>\<W>; \<K> \<subseteq> \<W>; | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 920 | \<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<W>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False\<rbrakk> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 921 | \<Longrightarrow> \<W> = \<K>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 922 | and ope: "\<forall>U\<in>\<K>. openin X U" and top: "topspace X \<subseteq> \<Union>\<K>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 923 | and non: "\<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<K>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 924 | unfolding \<A>_def by simp_all metis+ | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 925 | then obtain x where "x \<in> topspace X" "x \<notin> \<Union>(\<B> \<inter> \<K>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 926 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 927 | have "\<Union>(\<B> \<inter> \<K>) \<noteq> \<Union>\<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 928 | by (metis \<open>\<Union>\<B> = topspace X\<close> fin inf.bounded_iff non order_refl) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 929 | then have "\<exists>a. a \<notin> \<Union>(\<B> \<inter> \<K>) \<and> a \<in> \<Union>\<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 930 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 931 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 932 | using that by (metis U\<B>) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 933 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 934 | obtain C where C: "openin X C" "C \<in> \<K>" "x \<in> C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 935 | using \<open>x \<in> topspace X\<close> ope top by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 936 | then have "C \<subseteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 937 | by (metis openin_subset) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 938 | then have "(arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 939 | using openin_subbase C unfolding X [symmetric] by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 940 | moreover have "C \<noteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 941 | using \<open>\<K> \<in> \<A>\<close> \<open>C \<in> \<K>\<close> unfolding \<A>_def by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 942 | ultimately obtain \<V> W where W: "(finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to topspace X) W" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 943 | and "x \<in> W" "W \<in> \<V>" "\<Union>\<V> \<noteq> topspace X" "C = \<Union>\<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 944 | using C by (auto simp: union_of_def U\<B>) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 945 | then have "\<Union>\<V> \<subseteq> topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 946 | by (metis \<open>C \<subseteq> topspace X\<close>) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 947 | then have "topspace X \<notin> \<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 948 | using \<open>\<Union>\<V> \<noteq> topspace X\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 949 | then obtain \<B>' where \<B>': "finite \<B>'" "\<B>' \<subseteq> \<B>" "x \<in> \<Inter>\<B>'" "W = topspace X \<inter> \<Inter>\<B>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 950 | using W \<open>x \<in> W\<close> unfolding relative_to_def intersection_of_def by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 951 | then have "\<Inter>\<B>' \<subseteq> \<Union>\<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 952 | using \<open>W \<in> \<V>\<close> \<open>\<Union>\<V> \<noteq> topspace X\<close> \<open>\<Union>\<V> \<subseteq> topspace X\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 953 | then have "\<Inter>\<B>' \<subseteq> C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 954 | using U\<B> \<open>C = \<Union>\<V>\<close> \<open>W = topspace X \<inter> \<Inter>\<B>'\<close> \<open>W \<in> \<V>\<close> by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 955 | have "\<forall>b \<in> \<B>'. \<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 956 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 957 | fix b | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 958 | assume "b \<in> \<B>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 959 | have "insert b \<K> = \<K>" if neg: "\<not> (\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C'))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 960 | proof (rule *) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 961 | show "openin X W" if "W \<in> insert b \<K>" for W | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 962 | using that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 963 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 964 | have "b \<in> \<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 965 | using \<open>b \<in> \<B>'\<close> \<open>\<B>' \<subseteq> \<B>\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 966 | then have "\<exists>\<U>. finite \<U> \<and> \<U> \<subseteq> \<B> \<and> \<Inter>\<U> = b" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 967 |             by (rule_tac x="{b}" in exI) auto
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 968 | moreover have "\<Union>\<B> \<inter> b = b" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 969 | using \<B>'(2) \<open>b \<in> \<B>'\<close> by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 970 | ultimately show "openin X W" if "W = b" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 971 | using that \<open>b \<in> \<B>'\<close> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 972 | apply (simp add: openin_subbase flip: X) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 973 | apply (auto simp: arbitrary_def intersection_of_def relative_to_def intro!: union_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 974 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 975 | show "openin X W" if "W \<in> \<K>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 976 | by (simp add: \<open>W \<in> \<K>\<close> ope) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 977 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 978 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 979 | show "topspace X \<subseteq> \<Union> (insert b \<K>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 980 | using top by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 981 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 982 | show False if "finite \<F>" and "\<F> \<subseteq> insert b \<K>" "topspace X \<subseteq> \<Union>\<F>" for \<F> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 983 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 984 | have "insert b (\<F> \<inter> \<K>) = \<F>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 985 | using non that by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 986 | then show False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 987 | by (metis Int_lower2 finite_insert neg that(1) that(3)) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 988 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 989 | qed auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 990 | then show "\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 991 | using \<open>b \<in> \<B>'\<close> \<open>x \<notin> \<Union>(\<B> \<inter> \<K>)\<close> \<B>' | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 992 | by (metis IntI InterE Union_iff subsetD insertI1) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 993 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 994 | then obtain F where F: "\<forall>b \<in> \<B>'. finite (F b) \<and> F b \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b (F b))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 995 | by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 996 | let ?\<D> = "insert C (\<Union>(F ` \<B>'))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 997 | show False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 998 | proof (rule non) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 999 | have "topspace X \<subseteq> (\<Inter>b \<in> \<B>'. \<Union>(insert b (F b)))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1000 | using F by (simp add: INT_greatest) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1001 | also have "\<dots> \<subseteq> \<Union>?\<D>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1002 | using \<open>\<Inter>\<B>' \<subseteq> C\<close> by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1003 | finally show "topspace X \<subseteq> \<Union>?\<D>" . | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1004 | show "?\<D> \<subseteq> \<K>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1005 | using \<open>C \<in> \<K>\<close> F by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1006 | show "finite ?\<D>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1007 | using \<open>finite \<B>'\<close> F by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1008 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1009 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1010 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1011 | by (force simp: compact_space_def compactin_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1012 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1013 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1014 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1015 | corollary Alexander_subbase_alt: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1016 | assumes "U \<subseteq> \<Union>\<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1017 | and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; U \<subseteq> \<Union>C\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> U \<subseteq> \<Union>C'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1018 | and X: "topology | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1019 | (arbitrary union_of | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1020 | (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to U)) = X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1021 | shows "compact_space X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1022 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1023 | have "topspace X = U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1024 | using X topspace_subbase by fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1025 | have eq: "\<Union> (Collect ((\<lambda>x. x \<in> \<B>) relative_to U)) = U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1026 | unfolding relative_to_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1027 | using \<open>U \<subseteq> \<Union>\<B>\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1028 | have *: "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<C> \<and> \<Union>\<F> = topspace X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1029 | if "\<C> \<subseteq> Collect ((\<lambda>x. x \<in> \<B>) relative_to topspace X)" and UC: "\<Union>\<C> = topspace X" for \<C> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1030 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1031 | have "\<C> \<subseteq> (\<lambda>U. topspace X \<inter> U) ` \<B>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1032 | using that by (auto simp: relative_to_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1033 | then obtain \<B>' where "\<B>' \<subseteq> \<B>" and \<B>': "\<C> = (\<inter>) (topspace X) ` \<B>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1034 | by (auto simp: subset_image_iff) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1035 | moreover have "U \<subseteq> \<Union>\<B>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1036 | using \<B>' \<open>topspace X = U\<close> UC by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1037 | ultimately obtain \<C>' where "finite \<C>'" "\<C>' \<subseteq> \<B>'" "U \<subseteq> \<Union>\<C>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1038 | using fin [of \<B>'] \<open>topspace X = U\<close> \<open>U \<subseteq> \<Union>\<B>\<close> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1039 | then show ?thesis | 
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1040 | unfolding \<B>' ex_finite_subset_image \<open>topspace X = U\<close> by auto | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1041 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1042 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1043 | apply (rule Alexander_subbase [where \<B> = "Collect ((\<lambda>x. x \<in> \<B>) relative_to (topspace X))"]) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1044 | apply (simp flip: X) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1045 | apply (metis finite_intersection_of_relative_to eq) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1046 | apply (blast intro: *) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1047 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1048 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1049 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1050 | proposition continuous_map_componentwise: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1051 | "continuous_map X (product_topology Y I) f \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1052 | f ` (topspace X) \<subseteq> extensional I \<and> (\<forall>k \<in> I. continuous_map X (Y k) (\<lambda>x. f x k))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1053 | (is "?lhs \<longleftrightarrow> _ \<and> ?rhs") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1054 | proof (cases "\<forall>x \<in> topspace X. f x \<in> extensional I") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1055 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1056 | then have "f ` (topspace X) \<subseteq> extensional I" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1057 | by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1058 | moreover have ?rhs if L: ?lhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1059 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1060 |     have "openin X {x \<in> topspace X. f x k \<in> U}" if "k \<in> I" and "openin (Y k) U" for k U
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1061 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1062 |       have "openin (product_topology Y I) ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1063 | apply (simp add: openin_product_topology flip: arbitrary_union_of_relative_to) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1064 | apply (simp add: relative_to_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1065 | using that apply (blast intro: arbitrary_union_of_inc finite_intersection_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1066 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1067 |       with that have "openin X {x \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1068 | using L unfolding continuous_map_def by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1069 |       moreover have "{x \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))} = {x \<in> topspace X. f x k \<in> U}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1070 | using L by (auto simp: continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1071 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1072 | by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1073 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1074 | with that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1075 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1076 | by (auto simp: continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1077 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1078 | moreover have ?lhs if ?rhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1079 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1080 | have 1: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1081 | using that True by (auto simp: continuous_map_def PiE_iff) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1082 |     have 2: "{x \<in> S. \<exists>T\<in>\<T>. f x \<in> T} = (\<Union>T\<in>\<T>. {x \<in> S. f x \<in> T})" for S \<T>
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1083 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1084 |     have 3: "{x \<in> S. \<forall>U\<in>\<U>. f x \<in> U} = (\<Inter> (insert S ((\<lambda>U. {x \<in> S. f x \<in> U}) ` \<U>)))" for S \<U>
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1085 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1086 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1087 | unfolding continuous_map_def openin_product_topology arbitrary_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1088 | proof (clarsimp simp: all_union_of 1 2) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1089 | fix \<T> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1090 |       assume \<T>: "\<T> \<subseteq> Collect (finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U)
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1091 | relative_to (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1092 |       show "openin X (\<Union>T\<in>\<T>. {x \<in> topspace X. f x \<in> T})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1093 | proof (rule openin_Union; clarify) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1094 | fix S T | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1095 | assume "T \<in> \<T>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1096 | obtain \<U> where "T = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<U>" and "finite \<U>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1097 |           "\<U> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1098 | using subsetD [OF \<T> \<open>T \<in> \<T>\<close>] by (auto simp: intersection_of_def relative_to_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1099 |         with that show "openin X {x \<in> topspace X. f x \<in> T}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1100 | apply (simp add: continuous_map_def 1 cong: conj_cong) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1101 | unfolding 3 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1102 | apply (rule openin_Inter; auto) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1103 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1104 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1105 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1106 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1107 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1108 | by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1109 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1110 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1111 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1112 | by (auto simp: continuous_map_def PiE_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1113 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1114 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1115 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1116 | lemma continuous_map_componentwise_UNIV: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1117 | "continuous_map X (product_topology Y UNIV) f \<longleftrightarrow> (\<forall>k. continuous_map X (Y k) (\<lambda>x. f x k))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1118 | by (simp add: continuous_map_componentwise) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1119 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1120 | lemma continuous_map_product_projection [continuous_intros]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1121 | "k \<in> I \<Longrightarrow> continuous_map (product_topology X I) (X k) (\<lambda>x. x k)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1122 | using continuous_map_componentwise [of "product_topology X I" X I id] by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1123 | |
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1124 | declare continuous_map_from_subtopology [OF continuous_map_product_projection, continuous_intros] | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1125 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1126 | proposition open_map_product_projection: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1127 | assumes "i \<in> I" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1128 | shows "open_map (product_topology Y I) (Y i) (\<lambda>f. f i)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1129 | unfolding openin_product_topology all_union_of arbitrary_def open_map_def image_Union | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1130 | proof clarify | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1131 | fix \<V> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1132 | assume \<V>: "\<V> \<subseteq> Collect | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1133 | (finite intersection_of | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1134 |                 (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U) relative_to
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1135 | topspace (product_topology Y I))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1136 | show "openin (Y i) (\<Union>x\<in>\<V>. (\<lambda>f. f i) ` x)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1137 | proof (rule openin_Union, clarify) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1138 | fix S V | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1139 | assume "V \<in> \<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1140 | obtain \<F> where "finite \<F>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1141 | and V: "V = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1142 |       and \<F>: "\<F> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1143 | using subsetD [OF \<V> \<open>V \<in> \<V>\<close>] | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1144 | by (auto simp: intersection_of_def relative_to_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1145 | show "openin (Y i) ((\<lambda>f. f i) ` V)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1146 | proof (subst openin_subopen; clarify) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1147 | fix x f | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1148 | assume "f \<in> V" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1149 |       let ?T = "{a \<in> topspace(Y i).
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1150 | (\<lambda>j. if j = i then a | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1151 | else if j \<in> I then f j else undefined) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1152 | show "\<exists>T. openin (Y i) T \<and> f i \<in> T \<and> T \<subseteq> (\<lambda>f. f i) ` V" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1153 | proof (intro exI conjI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1154 | show "openin (Y i) ?T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1155 | proof (rule openin_continuous_map_preimage) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1156 | have "continuous_map (Y i) (Y k) (\<lambda>x. if k = i then x else f k)" if "k \<in> I" for k | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1157 | proof (cases "k=i") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1158 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1159 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1160 | by (metis (mono_tags) continuous_map_id eq_id_iff) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1161 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1162 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1163 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1164 | by simp (metis IntD1 PiE_iff V \<open>f \<in> V\<close> that) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1165 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1166 | then show "continuous_map (Y i) (product_topology Y I) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1167 | (\<lambda>x j. if j = i then x else if j \<in> I then f j else undefined)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1168 | by (auto simp: continuous_map_componentwise assms extensional_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1169 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1170 | have "openin (product_topology Y I) (\<Pi>\<^sub>E i\<in>I. topspace (Y i))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1171 | by (metis openin_topspace topspace_product_topology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1172 | moreover have "openin (product_topology Y I) (\<Inter>B\<in>\<F>. (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> B)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1173 |                          if "\<F> \<noteq> {}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1174 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1175 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1176 | proof (rule openin_Inter) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1177 | show "\<And>X. X \<in> (\<inter>) (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) ` \<F> \<Longrightarrow> openin (product_topology Y I) X" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1178 | unfolding openin_product_topology relative_to_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1179 | apply (clarify intro!: arbitrary_union_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1180 | apply (rename_tac F) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1181 | apply (rule_tac x=F in exI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1182 | using subsetD [OF \<F>] | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1183 | apply (force intro: finite_intersection_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1184 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1185 |             qed (use \<open>finite \<F>\<close> \<open>\<F> \<noteq> {}\<close> in auto)
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1186 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1187 | ultimately show "openin (product_topology Y I) ((\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1188 | by (auto simp only: Int_Inter_eq split: if_split) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1189 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1190 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1191 | have eqf: "(\<lambda>j. if j = i then f i else if j \<in> I then f j else undefined) = f" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1192 | using PiE_arb V \<open>f \<in> V\<close> by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1193 | show "f i \<in> ?T" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1194 | using V assms \<open>f \<in> V\<close> by (auto simp: PiE_iff eqf) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1195 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1196 | show "?T \<subseteq> (\<lambda>f. f i) ` V" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1197 | unfolding V by (auto simp: intro!: rev_image_eqI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1198 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1199 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1200 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1201 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1202 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1203 | lemma retraction_map_product_projection: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1204 | assumes "i \<in> I" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1205 | shows "(retraction_map (product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1206 |          (topspace (product_topology X I) = {}) \<longrightarrow> topspace (X i) = {})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1207 | (is "?lhs = ?rhs") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1208 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1209 | assume ?lhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1210 | then show ?rhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1211 | using retraction_imp_surjective_map by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1212 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1213 | assume R: ?rhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1214 | show ?lhs | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1215 |   proof (cases "topspace (product_topology X I) = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1216 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1217 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1218 | using R by (auto simp: retraction_map_def retraction_maps_def continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1219 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1220 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1221 | have *: "\<exists>g. continuous_map (X i) (product_topology X I) g \<and> (\<forall>x\<in>topspace (X i). g x i = x)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1222 | if z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" for z | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1223 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1224 | have cm: "continuous_map (X i) (X j) (\<lambda>x. if j = i then x else z j)" if "j \<in> I" for j | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1225 | using \<open>j \<in> I\<close> z by (case_tac "j = i") auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1226 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1227 | using \<open>i \<in> I\<close> that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1228 | by (rule_tac x="\<lambda>x j. if j = i then x else z j" in exI) (auto simp: continuous_map_componentwise PiE_iff extensional_def cm) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1229 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1230 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1231 | using \<open>i \<in> I\<close> False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1232 | by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1233 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1234 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 1235 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1236 | subsection \<open>Open Pi-sets in the product topology\<close> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1237 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1238 | proposition openin_PiE_gen: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1239 | "openin (product_topology X I) (PiE I S) \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1240 |         PiE I S = {} \<or>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1241 |         finite {i \<in> I. ~(S i = topspace(X i))} \<and> (\<forall>i \<in> I. openin (X i) (S i))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1242 | (is "?lhs \<longleftrightarrow> _ \<or> ?rhs") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1243 | proof (cases "PiE I S = {}")
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1244 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1245 | moreover have "?lhs = ?rhs" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1246 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1247 | assume L: ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1248 | moreover | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1249 | obtain z where z: "z \<in> PiE I S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1250 | using False by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1251 |     ultimately obtain U where fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1252 |       and "Pi\<^sub>E I U \<noteq> {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1253 | and sub: "Pi\<^sub>E I U \<subseteq> Pi\<^sub>E I S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1254 | by (fastforce simp add: openin_product_topology_alt) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1255 | then have *: "\<And>i. i \<in> I \<Longrightarrow> U i \<subseteq> S i" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1256 | by (simp add: subset_PiE) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1257 | show ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1258 | proof (intro conjI ballI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1259 |       show "finite {i \<in> I. S i \<noteq> topspace (X i)}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1260 | apply (rule finite_subset [OF _ fin], clarify) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1261 | using * | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1262 | by (metis False L openin_subset topspace_product_topology subset_PiE subset_antisym) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1263 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1264 | fix i :: "'a" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1265 | assume "i \<in> I" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1266 | then show "openin (X i) (S i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1267 | using open_map_product_projection [of i I X] L | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1268 | apply (simp add: open_map_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1269 | apply (drule_tac x="PiE I S" in spec) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1270 | apply (simp add: False image_projection_PiE split: if_split_asm) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1271 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1272 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1273 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1274 | assume ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1275 | then show ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1276 | apply (simp only: openin_product_topology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1277 | apply (rule arbitrary_union_of_inc) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1278 | apply (auto simp: product_topology_base_alt) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1279 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1280 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1281 | ultimately show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1282 | by simp | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1283 | qed simp | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1284 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1285 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1286 | corollary openin_PiE: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1287 |    "finite I \<Longrightarrow> openin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. openin (X i) (S i))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1288 | by (simp add: openin_PiE_gen) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1289 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1290 | proposition compact_space_product_topology: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1291 | "compact_space(product_topology X I) \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1292 |     topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. compact_space(X i))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1293 | (is "?lhs = ?rhs") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1294 | proof (cases "topspace(product_topology X I) = {}")
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1295 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1296 | then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1297 | by auto | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1298 | show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1299 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1300 | assume L: ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1301 | show ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1302 | proof (clarsimp simp add: False compact_space_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1303 | fix i | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1304 | assume "i \<in> I" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1305 | with L have "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1306 | by (simp add: continuous_map_product_projection) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1307 | moreover | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1308 | have "\<And>x. x \<in> topspace (X i) \<Longrightarrow> x \<in> (\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1309 | using \<open>i \<in> I\<close> z | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1310 | apply (rule_tac x="\<lambda>j. if j = i then x else if j \<in> I then z j else undefined" in image_eqI, auto) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1311 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1312 | then have "(\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = topspace (X i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1313 | using \<open>i \<in> I\<close> z by auto | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1314 | ultimately show "compactin (X i) (topspace (X i))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1315 | by (metis L compact_space_def image_compactin topspace_product_topology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1316 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1317 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1318 | assume R: ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1319 | show ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1320 |     proof (cases "I = {}")
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1321 | case True | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1322 | with R show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1323 | by (simp add: compact_space_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1324 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1325 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1326 | then obtain i where "i \<in> I" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1327 | by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1328 | show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1329 | using R | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1330 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1331 | assume com [rule_format]: "\<forall>i\<in>I. compact_space (X i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1332 |         let ?\<C> = "{{f. f i \<in> U} |i U. i \<in> I \<and> openin (X i) U}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1333 | show "compact_space (product_topology X I)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1334 | proof (rule Alexander_subbase_alt) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1335 | show "topspace (product_topology X I) \<subseteq> \<Union>?\<C>" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1336 | unfolding topspace_product_topology using \<open>i \<in> I\<close> by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1337 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1338 | fix C | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1339 | assume Csub: "C \<subseteq> ?\<C>" and UC: "topspace (product_topology X I) \<subseteq> \<Union>C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1340 |           define \<D> where "\<D> \<equiv> \<lambda>i. {U. openin (X i) U \<and> {f. f i \<in> U} \<in> C}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1341 | show "\<exists>C'. finite C' \<and> C' \<subseteq> C \<and> topspace (product_topology X I) \<subseteq> \<Union>C'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1342 | proof (cases "\<exists>i. i \<in> I \<and> topspace (X i) \<subseteq> \<Union>(\<D> i)") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1343 | case True | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1344 | then obtain i where "i \<in> I" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1345 | and i: "topspace (X i) \<subseteq> \<Union>(\<D> i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1346 | unfolding \<D>_def by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1347 | then have *: "\<And>\<U>. \<lbrakk>Ball \<U> (openin (X i)); topspace (X i) \<subseteq> \<Union>\<U>\<rbrakk> \<Longrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1348 | \<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace (X i) \<subseteq> \<Union>\<F>" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1349 | using com [OF \<open>i \<in> I\<close>] by (auto simp: compact_space_def compactin_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1350 | have "topspace (X i) \<subseteq> \<Union>(\<D> i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1351 | using i by auto | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1352 | with * obtain \<F> where "finite \<F> \<and> \<F> \<subseteq> (\<D> i) \<and> topspace (X i) \<subseteq> \<Union>\<F>" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1353 | unfolding \<D>_def by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1354 | with \<open>i \<in> I\<close> show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1355 | unfolding \<D>_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1356 |               by (rule_tac x="(\<lambda>U. {x. x i \<in> U}) ` \<F>" in exI) auto
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1357 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1358 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1359 | then have "\<forall>i \<in> I. \<exists>y. y \<in> topspace (X i) \<and> y \<notin> \<Union>(\<D> i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1360 | by force | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1361 | then obtain g where g: "\<And>i. i \<in> I \<Longrightarrow> g i \<in> topspace (X i) \<and> g i \<notin> \<Union>(\<D> i)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1362 | by metis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1363 | then have "(\<lambda>i. if i \<in> I then g i else undefined) \<in> topspace (product_topology X I)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1364 | by (simp add: PiE_I) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1365 | moreover have "(\<lambda>i. if i \<in> I then g i else undefined) \<notin> \<Union>C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1366 | using Csub g unfolding \<D>_def by force | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1367 | ultimately show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1368 | using UC by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1369 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1370 | qed (simp add: product_topology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1371 | qed (simp add: compact_space_topspace_empty) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1372 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1373 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1374 | qed (simp add: compact_space_topspace_empty) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1375 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1376 | corollary compactin_PiE: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1377 | "compactin (product_topology X I) (PiE I S) \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1378 |         PiE I S = {} \<or> (\<forall>i \<in> I. compactin (X i) (S i))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1379 | by (auto simp: compactin_subspace subtopology_PiE subset_PiE compact_space_product_topology | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1380 | PiE_eq_empty_iff) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1381 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1382 | lemma in_product_topology_closure_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1383 | "z \<in> (product_topology X I) closure_of S | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1384 | \<Longrightarrow> i \<in> I \<Longrightarrow> z i \<in> ((X i) closure_of ((\<lambda>x. x i) ` S))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1385 | using continuous_map_product_projection | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1386 | by (force simp: continuous_map_eq_image_closure_subset image_subset_iff) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1387 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1388 | lemma homeomorphic_space_singleton_product: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1389 |    "product_topology X {k} homeomorphic_space (X k)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1390 | unfolding homeomorphic_space | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1391 | apply (rule_tac x="\<lambda>x. x k" in exI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1392 | apply (rule bijective_open_imp_homeomorphic_map) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1393 | apply (simp_all add: continuous_map_product_projection open_map_product_projection) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1394 | unfolding PiE_over_singleton_iff | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1395 | apply (auto simp: image_iff inj_on_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1396 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1397 | |
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1398 | subsection\<open>Relationship with connected spaces, paths, etc.\<close> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1399 | |
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1400 | proposition connected_space_product_topology: | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1401 | "connected_space(product_topology X I) \<longleftrightarrow> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1402 |     (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = {} \<or> (\<forall>i \<in> I. connected_space(X i))"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1403 | (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1404 | proof (cases ?eq) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1405 | case False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1406 | moreover have "?lhs = ?rhs" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1407 | proof | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1408 | assume ?lhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1409 | moreover | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1410 | have "connectedin(X i) (topspace(X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1411 | if "i \<in> I" and ci: "connectedin(product_topology X I) (topspace(product_topology X I))" for i | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1412 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1413 | have cm: "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1414 | by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1415 | show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1416 | using connectedin_continuous_map_image [OF cm ci] \<open>i \<in> I\<close> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1417 | by (simp add: False image_projection_PiE) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1418 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1419 | ultimately show ?rhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1420 | by (meson connectedin_topspace) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1421 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1422 | assume cs [rule_format]: ?rhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1423 | have False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1424 |       if disj: "U \<inter> V = {}" and subUV: "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<subseteq> U \<union> V"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1425 | and U: "openin (product_topology X I) U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1426 | and V: "openin (product_topology X I) V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1427 |         and "U \<noteq> {}" "V \<noteq> {}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1428 | for U V | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1429 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1430 | obtain f where "f \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1431 |         using \<open>U \<noteq> {}\<close> by blast
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1432 | then have f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1433 | using U openin_subset by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1434 | have "U \<subseteq> topspace(product_topology X I)" "V \<subseteq> topspace(product_topology X I)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1435 | using U V openin_subset by blast+ | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1436 | moreover have "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<subseteq> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1437 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1438 |         obtain C where "(finite intersection_of (\<lambda>F. \<exists>i U. F = {x. x i \<in> U} \<and> i \<in> I \<and> openin (X i) U) relative_to
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1439 | (\<Pi>\<^sub>E i\<in>I. topspace (X i))) C" "C \<subseteq> U" "f \<in> C" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1440 | using U \<open>f \<in> U\<close> unfolding openin_product_topology union_of_def by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1441 | then obtain \<T> where "finite \<T>" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1442 |           and t: "\<And>C. C \<in> \<T> \<Longrightarrow> \<exists>i u. (i \<in> I \<and> openin (X i) u) \<and> C = {x. x i \<in> u}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1443 | and subU: "topspace (product_topology X I) \<inter> \<Inter>\<T> \<subseteq> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1444 | and ftop: "f \<in> topspace (product_topology X I)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1445 | and fint: "f \<in> \<Inter> \<T>" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1446 | by (fastforce simp: relative_to_def intersection_of_def subset_iff) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1447 |         let ?L = "\<Union>C\<in>\<T>. {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1448 | obtain L where "finite L" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1449 |           and L: "\<And>i U. \<lbrakk>i \<in> I; openin (X i) U; U \<subset> topspace(X i); {x. x i \<in> U} \<in> \<T>\<rbrakk> \<Longrightarrow> i \<in> L"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1450 | proof | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1451 | show "finite ?L" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1452 | proof (rule finite_Union) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1453 | fix M | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1454 |             assume "M \<in> (\<lambda>C. {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}) ` \<T>"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1455 |             then obtain C where "C \<in> \<T>" and C: "M = {i. (\<lambda>x. x i) ` C \<subset> topspace (X i)}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1456 | by blast | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1457 |             then obtain j V where "j \<in> I" and ope: "openin (X j) V" and Ceq: "C = {x. x j \<in> V}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1458 | using t by meson | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1459 | then have "f j \<in> V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1460 | using \<open>C \<in> \<T>\<close> fint by force | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1461 |             then have "(\<lambda>x. x k) ` {x. x j \<in> V} = UNIV" if "k \<noteq> j" for k
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1462 | using that | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1463 | apply (clarsimp simp add: set_eq_iff) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1464 | apply (rule_tac x="\<lambda>m. if m = k then x else f m" in image_eqI, auto) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1465 | done | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1466 |             then have "{i. (\<lambda>x. x i) ` C \<subset> topspace (X i)} \<subseteq> {j}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1467 | using Ceq by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1468 | then show "finite M" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1469 | using C finite_subset by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1470 | qed (use \<open>finite \<T>\<close> in blast) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1471 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1472 | fix i U | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1473 |           assume  "i \<in> I" and ope: "openin (X i) U" and psub: "U \<subset> topspace (X i)" and int: "{x. x i \<in> U} \<in> \<T>"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1474 | then show "i \<in> ?L" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1475 |             by (rule_tac a="{x. x i \<in> U}" in UN_I) (force+)
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1476 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1477 | show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1478 | proof | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1479 | fix h | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1480 | assume h: "h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1481 | define g where "g \<equiv> \<lambda>i. if i \<in> L then f i else h i" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1482 | have gin: "g \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1483 | unfolding g_def using f h by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1484 | moreover have "g \<in> X" if "X \<in> \<T>" for X | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1485 | using fint openin_subset t [OF that] L g_def h that by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1486 | ultimately have "g \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1487 | using subU by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1488 |           have "h \<in> U" if "finite M" "h \<in> PiE I (topspace \<circ> X)" "{i \<in> I. h i \<noteq> g i} \<subseteq> M" for M h
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1489 | using that | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1490 | proof (induction arbitrary: h) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1491 | case empty | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1492 | then show ?case | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1493 | using PiE_ext \<open>g \<in> U\<close> gin by force | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1494 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1495 | case (insert i M) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1496 | define f where "f \<equiv> \<lambda>j. if j = i then g i else h j" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1497 | have fin: "f \<in> PiE I (topspace \<circ> X)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1498 | unfolding f_def using gin insert.prems(1) by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1499 |             have subM: "{j \<in> I. f j \<noteq> g j} \<subseteq> M"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1500 | unfolding f_def using insert.prems(2) by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1501 | have "f \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1502 | using insert.IH [OF fin subM] . | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1503 | show ?case | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1504 | proof (cases "h \<in> V") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1505 | case True | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1506 | show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1507 | proof (cases "i \<in> I") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1508 | case True | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1509 |                 let ?U = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> U}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1510 |                 let ?V = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> V}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1511 | have False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1512 | proof (rule connected_spaceD [OF cs [OF \<open>i \<in> I\<close>]]) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1513 | have "\<And>k. k \<in> I \<Longrightarrow> continuous_map (X i) (X k) (\<lambda>x. if k = i then x else h k)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1514 | using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1515 | then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x j. if j = i then x else h j)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1516 | using \<open>i \<in> I\<close> insert.prems(1) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1517 | by (auto simp: continuous_map_componentwise extensional_def) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1518 | show "openin (X i) ?U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1519 | by (rule openin_continuous_map_preimage [OF cm U]) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1520 | show "openin (X i) ?V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1521 | by (rule openin_continuous_map_preimage [OF cm V]) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1522 | show "topspace (X i) \<subseteq> ?U \<union> ?V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1523 | proof clarsimp | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1524 | fix x | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1525 | assume "x \<in> topspace (X i)" and "(\<lambda>j. if j = i then x else h j) \<notin> V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1526 | with True subUV \<open>h \<in> Pi\<^sub>E I (topspace \<circ> X)\<close> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1527 | show "(\<lambda>j. if j = i then x else h j) \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1528 | by (drule_tac c="(\<lambda>j. if j = i then x else h j)" in subsetD) auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1529 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1530 |                   show "?U \<inter> ?V = {}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1531 | using disj by blast | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1532 |                   show "?U \<noteq> {}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1533 |                     using \<open>U \<noteq> {}\<close> f_def
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1534 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1535 | have "(\<lambda>j. if j = i then g i else h j) \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1536 | using \<open>f \<in> U\<close> f_def by blast | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1537 | moreover have "f i \<in> topspace (X i)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1538 | by (metis PiE_iff True comp_apply fin) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1539 | ultimately have "\<exists>b. b \<in> topspace (X i) \<and> (\<lambda>a. if a = i then b else h a) \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1540 | using f_def by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1541 | then show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1542 | by blast | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1543 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1544 | have "(\<lambda>j. if j = i then h i else h j) = h" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1545 | by force | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1546 | moreover have "h i \<in> topspace (X i)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1547 | using True insert.prems(1) by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1548 |                   ultimately show "?V \<noteq> {}"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1549 | using \<open>h \<in> V\<close> by force | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1550 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1551 | then show ?thesis .. | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1552 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1553 | case False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1554 | show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1555 | proof (cases "h = f") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1556 | case True | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1557 | show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1558 | by (rule insert.IH [OF insert.prems(1)]) (simp add: True subM) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1559 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1560 | case False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1561 | then show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1562 | using gin insert.prems(1) \<open>i \<notin> I\<close> unfolding f_def by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1563 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1564 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1565 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1566 | case False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1567 | then show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1568 | using subUV insert.prems(1) by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1569 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1570 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1571 | then show "h \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1572 | unfolding g_def using PiE_iff \<open>finite L\<close> h by fastforce | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1573 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1574 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1575 | ultimately show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1576 |         using disj inf_absorb2 \<open>V \<noteq> {}\<close> by fastforce
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1577 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1578 | then show ?lhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1579 | unfolding connected_space_def | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1580 | by auto | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1581 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1582 | ultimately show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1583 | by simp | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1584 | qed (simp add: connected_space_topspace_empty) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1585 | |
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1586 | |
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1587 | lemma connectedin_PiE: | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1588 | "connectedin (product_topology X I) (PiE I S) \<longleftrightarrow> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1589 |         PiE I S = {} \<or> (\<forall>i \<in> I. connectedin (X i) (S i))"
 | 
| 71172 | 1590 | by (fastforce simp add: connectedin_def subtopology_PiE connected_space_product_topology subset_PiE PiE_eq_empty_iff) | 
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1591 | |
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1592 | lemma path_connected_space_product_topology: | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1593 | "path_connected_space(product_topology X I) \<longleftrightarrow> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1594 |     topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. path_connected_space(X i))"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1595 | (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1596 | proof (cases ?eq) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1597 | case False | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1598 | moreover have "?lhs = ?rhs" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1599 | proof | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1600 | assume L: ?lhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1601 | show ?rhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1602 | proof (clarsimp simp flip: path_connectedin_topspace) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1603 | fix i :: "'a" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1604 | assume "i \<in> I" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1605 | have cm: "continuous_map (product_topology X I) (X i) (\<lambda>f. f i)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1606 | by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1607 | show "path_connectedin (X i) (topspace (X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1608 | using path_connectedin_continuous_map_image [OF cm L [unfolded path_connectedin_topspace [symmetric]]] | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1609 | by (metis \<open>i \<in> I\<close> False retraction_imp_surjective_map retraction_map_product_projection) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1610 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1611 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1612 | assume R [rule_format]: ?rhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1613 | show ?lhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1614 | unfolding path_connected_space_def topspace_product_topology | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1615 | proof clarify | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1616 | fix x y | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1617 | assume x: "x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and y: "y \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1618 | have "\<forall>i. \<exists>g. i \<in> I \<longrightarrow> pathin (X i) g \<and> g 0 = x i \<and> g 1 = y i" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1619 | using PiE_mem R path_connected_space_def x y by force | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1620 | then obtain g where g: "\<And>i. i \<in> I \<Longrightarrow> pathin (X i) (g i) \<and> g i 0 = x i \<and> g i 1 = y i" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1621 | by metis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1622 | with x y show "\<exists>g. pathin (product_topology X I) g \<and> g 0 = x \<and> g 1 = y" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1623 | apply (rule_tac x="\<lambda>a. \<lambda>i \<in> I. g i a" in exI) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1624 | apply (force simp: pathin_def continuous_map_componentwise) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1625 | done | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1626 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1627 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1628 | ultimately show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1629 | by simp | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1630 | qed (simp add: path_connected_space_topspace_empty) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1631 | |
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1632 | lemma path_connectedin_PiE: | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1633 | "path_connectedin (product_topology X I) (PiE I S) \<longleftrightarrow> | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1634 |     PiE I S = {} \<or> (\<forall>i \<in> I. path_connectedin (X i) (S i))"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1635 | by (fastforce simp add: path_connectedin_def subtopology_PiE path_connected_space_product_topology subset_PiE PiE_eq_empty_iff topspace_subtopology_subset) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1636 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1637 | subsection \<open>Projections from a function topology to a component\<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1638 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1639 | lemma quotient_map_product_projection: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1640 | assumes "i \<in> I" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1641 | shows "quotient_map(product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1642 |            (topspace(product_topology X I) = {} \<longrightarrow> topspace(X i) = {})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1643 | (is "?lhs = ?rhs") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1644 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1645 | assume ?lhs with assms show ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1646 | by (auto simp: continuous_open_quotient_map open_map_product_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1647 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1648 | assume ?rhs with assms show ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1649 | by (auto simp: Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1650 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1651 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1652 | lemma product_topology_homeomorphic_component: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1653 |   assumes "i \<in> I" "\<And>j. \<lbrakk>j \<in> I; j \<noteq> i\<rbrakk> \<Longrightarrow> \<exists>a. topspace(X j) = {a}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1654 | shows "product_topology X I homeomorphic_space (X i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1655 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1656 | have "quotient_map (product_topology X I) (X i) (\<lambda>x. x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1657 | using assms by (force simp add: quotient_map_product_projection PiE_eq_empty_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1658 | moreover | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1659 | have "inj_on (\<lambda>x. x i) (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1660 | using assms by (auto simp: inj_on_def PiE_iff) (metis extensionalityI singletonD) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1661 | ultimately show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1662 | unfolding homeomorphic_space_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1663 | by (rule_tac x="\<lambda>x. x i" in exI) (simp add: homeomorphic_map_def flip: homeomorphic_map_maps) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1664 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1665 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1666 | lemma topological_property_of_product_component: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1667 | assumes major: "P (product_topology X I)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1668 | and minor: "\<And>z i. \<lbrakk>z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i)); P(product_topology X I); i \<in> I\<rbrakk> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1669 |                       \<Longrightarrow> P(subtopology (product_topology X I) (PiE I (\<lambda>j. if j = i then topspace(X i) else {z j})))"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1670 | (is "\<And>z i. \<lbrakk>_; _; _\<rbrakk> \<Longrightarrow> P (?SX z i)") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1671 | and PQ: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1672 |   shows "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. Q(X i))"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1673 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1674 |   have "Q(X i)" if "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<noteq> {}" "i \<in> I" for i
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1675 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1676 | from that obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1677 | by force | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1678 | have "?SX f i homeomorphic_space X i" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1679 | apply (simp add: subtopology_PiE ) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1680 |       using product_topology_homeomorphic_component [OF \<open>i \<in> I\<close>, of "\<lambda>j. subtopology (X j) (if j = i then topspace (X i) else {f j})"]
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1681 | using f by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1682 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1683 | using minor [OF f major \<open>i \<in> I\<close>] PQ by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1684 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1685 | then show ?thesis by metis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 1686 | qed | 
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 1687 | |
| 64289 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 hoelzl parents: diff
changeset | 1688 | end |