author | wenzelm |
Mon, 20 Feb 2023 21:40:52 +0100 | |
changeset 77319 | 87698fe320bb |
parent 76838 | 04c7ec38874e |
child 78200 | 264f2b69d09c |
permissions | -rw-r--r-- |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
71172
diff
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:
71172
diff
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:
71172
diff
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:
71172
diff
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:
69922
diff
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:
69918
diff
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:
69922
diff
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:
69922
diff
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:
69922
diff
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:
69922
diff
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:
69945
diff
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:
69945
diff
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:
69945
diff
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:
68833
diff
changeset
|
90 |
proposition product_topology: |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
91 |
"product_topology X I = |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
92 |
topology |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
93 |
(arbitrary union_of |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
94 |
((finite intersection_of |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
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:
68833
diff
changeset
|
97 |
(is "_ = topology (_ union_of ((_ intersection_of ?\<Psi>) relative_to ?TOP))") |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
98 |
proof - |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
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:
68833
diff
changeset
|
101 |
proof - |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
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:
68833
diff
changeset
|
104 |
proof - |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
changeset
|
106 |
using \<U> by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
changeset
|
108 |
by metis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
109 |
obtain U where "U \<in> \<U>" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
110 |
using \<open>\<U> \<noteq> {}\<close> by blast |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
changeset
|
112 |
show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
113 |
proof (intro conjI exI) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
114 |
show "finite (\<Union>U\<in>\<U>. ?F U)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
115 |
using Y \<open>finite' \<U>\<close> by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
changeset
|
117 |
proof |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
118 |
have *: "f \<in> U" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
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:
68833
diff
changeset
|
120 |
and "\<forall>i\<in>I. f i \<in> topspace (X i)" and "f \<in> extensional I" for f U |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
121 |
by (smt (verit) PiE_iff Y that) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
122 |
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:
68833
diff
changeset
|
123 |
by (auto simp: PiE_iff *) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
124 |
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:
68833
diff
changeset
|
125 |
using Y openin_subset \<open>finite' \<U>\<close> by fastforce |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
126 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
127 |
qed (use Y openin_subset in \<open>blast+\<close>) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
128 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
129 |
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:
68833
diff
changeset
|
130 |
if \<U>: "\<U> \<subseteq> Collect ?\<Psi>" and "finite \<U>" for \<U> |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
131 |
proof (cases "\<U>={}") |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
132 |
case True |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
133 |
then show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
134 |
apply (rule_tac x="{?TOP}" in exI, simp) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
135 |
apply (rule_tac x="\<lambda>i. topspace (X i)" in exI) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
136 |
apply force |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
137 |
done |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
138 |
next |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
139 |
case False |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
140 |
then obtain U where "U \<in> \<U>" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
141 |
by blast |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
142 |
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:
68833
diff
changeset
|
143 |
using \<U> by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
144 |
then obtain J Y where |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
145 |
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:
68833
diff
changeset
|
146 |
by metis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
147 |
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:
68833
diff
changeset
|
148 |
show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
149 |
proof (intro conjI exI) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
150 |
show "finite (?G ` \<U>)" "?G ` \<U> \<noteq> {}" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
151 |
using \<open>finite \<U>\<close> \<open>U \<in> \<U>\<close> by blast+ |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
152 |
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:
68833
diff
changeset
|
153 |
using Y by force |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
154 |
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:
68833
diff
changeset
|
155 |
apply clarsimp |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
156 |
apply (rule_tac x= "(\<lambda>i. if i = J U then Y U else topspace (X i))" in exI) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
157 |
apply (auto simp: *) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
158 |
done |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
159 |
next |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
160 |
show "(\<Inter>U\<in>\<U>. ?G U) = ?TOP \<inter> \<Inter>\<U>" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
161 |
proof |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
162 |
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))" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
163 |
by (simp add: PiE_mono Y \<open>U \<in> \<U>\<close> openin_subset) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
164 |
then have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
165 |
using \<open>U \<in> \<U>\<close> by fastforce |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
166 |
moreover have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> \<Inter>\<U>" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
167 |
using PiE_mem Y by fastforce |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
168 |
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:
68833
diff
changeset
|
169 |
by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
170 |
qed (use Y in fastforce) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
171 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
172 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
173 |
show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
174 |
unfolding relative_to_def intersection_of_def |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
175 |
by (safe; blast dest!: 1 2) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
176 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
177 |
show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
178 |
unfolding product_topology_def generate_topology_on_eq |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
179 |
apply (rule arg_cong [where f = topology]) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
180 |
apply (rule arg_cong [where f = "(union_of)arbitrary"]) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
181 |
apply (force simp: *) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
182 |
done |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
183 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
184 |
|
69874 | 185 |
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
|
186 |
"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:
69680
diff
changeset
|
187 |
proof |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
188 |
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:
68833
diff
changeset
|
189 |
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
|
190 |
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
|
191 |
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
|
192 |
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
|
193 |
then show "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<subseteq> topspace (product_topology T I)" |
71633 | 194 |
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
|
195 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
196 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
197 |
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
|
198 |
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
|
199 |
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:
68833
diff
changeset
|
200 |
unfolding product_topology_def |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
201 |
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
|
202 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
203 |
proposition product_topology_open_contains_basis: |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
204 |
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
|
205 |
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:
69680
diff
changeset
|
206 |
proof - |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
207 |
define IT where "IT \<equiv> \<lambda>X. {i. X i \<noteq> topspace (T i)}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
208 |
have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite (IT X)} U" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
209 |
using assms unfolding product_topology_def IT_def by (intro openin_topology_generated_by) auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
210 |
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 (IT X) \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U" |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
211 |
proof induction |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
212 |
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
|
213 |
then obtain XU XV where H: |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
214 |
"x \<in> Pi\<^sub>E I XU" "\<And>i. openin (T i) (XU i)" "finite (IT XU)" "Pi\<^sub>E I XU \<subseteq> U" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
215 |
"x \<in> Pi\<^sub>E I XV" "\<And>i. openin (T i) (XV i)" "finite (IT XV)" "Pi\<^sub>E I XV \<subseteq> V" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
216 |
by (meson Int_iff) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
217 |
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
|
218 |
have "Pi\<^sub>E I X \<subseteq> Pi\<^sub>E I XU \<inter> Pi\<^sub>E I XV" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
219 |
by (auto simp add: PiE_iff X_def) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
220 |
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
|
221 |
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
|
222 |
unfolding X_def using H by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
223 |
moreover have "finite (IT X)" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
224 |
apply (rule rev_finite_subset[of "IT XU \<union> IT XV"]) |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
225 |
using H by (auto simp: X_def IT_def) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
226 |
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
|
227 |
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
|
228 |
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
|
229 |
by auto |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
230 |
next |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
231 |
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
|
232 |
then obtain k where "k \<in> K" "x \<in> k" by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
233 |
with \<open>k \<in> K\<close> UN show ?case |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
234 |
by (meson Sup_upper2) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
235 |
qed auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
236 |
then show ?thesis using \<open>x \<in> U\<close> IT_def by blast |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
237 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
238 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
239 |
lemma product_topology_empty_discrete: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
240 |
"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:
69918
diff
changeset
|
241 |
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
|
242 |
|
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
243 |
lemma openin_product_topology: |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
244 |
"openin (product_topology X I) = |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
245 |
arbitrary union_of |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
246 |
((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:
68833
diff
changeset
|
247 |
relative_to topspace (product_topology X I))" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
248 |
by (simp add: istopology_subbase product_topology) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
249 |
|
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
250 |
lemma subtopology_PiE: |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
251 |
"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:
68833
diff
changeset
|
252 |
proof - |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
253 |
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:
68833
diff
changeset
|
254 |
let ?X = "\<Pi>\<^sub>E i\<in>I. topspace (X i)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
255 |
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:
68833
diff
changeset
|
256 |
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:
68833
diff
changeset
|
257 |
by (rule finite_intersection_of_relative_to) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
258 |
also have "\<dots> = finite intersection_of |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
259 |
((\<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:
68833
diff
changeset
|
260 |
relative_to ?X \<inter> Pi\<^sub>E I S) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
261 |
relative_to ?X \<inter> Pi\<^sub>E I S" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
262 |
apply (rule arg_cong2 [where f = "(relative_to)"]) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
263 |
apply (rule arg_cong [where f = "(intersection_of)finite"]) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
264 |
apply (rule ext) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
265 |
apply (auto simp: relative_to_def intersection_of_def) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
266 |
done |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
267 |
finally |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
268 |
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:
68833
diff
changeset
|
269 |
finite intersection_of |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
270 |
(\<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:
68833
diff
changeset
|
271 |
relative_to ?X \<inter> Pi\<^sub>E I S" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
272 |
by (metis finite_intersection_of_relative_to) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
273 |
then show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
274 |
unfolding topology_eq |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
275 |
apply clarify |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
276 |
apply (simp add: openin_product_topology flip: openin_relative_to) |
71172 | 277 |
apply (simp add: arbitrary_union_of_relative_to flip: PiE_Int) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
278 |
done |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
279 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
280 |
|
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
281 |
lemma product_topology_base_alt: |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
282 |
"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:
69922
diff
changeset
|
283 |
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
284 |
(\<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:
68833
diff
changeset
|
285 |
(is "?lhs = ?rhs") |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
286 |
proof - |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
287 |
have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
288 |
unfolding all_relative_to all_intersection_of topspace_product_topology |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
289 |
proof clarify |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
290 |
fix \<F> |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
291 |
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:
68833
diff
changeset
|
292 |
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:
68833
diff
changeset
|
293 |
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:
68833
diff
changeset
|
294 |
proof (induction) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
295 |
case (insert F \<F>) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
296 |
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:
68833
diff
changeset
|
297 |
and fin: "finite {i \<in> I. U i \<noteq> topspace (X i)}" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
298 |
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:
68833
diff
changeset
|
299 |
by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
300 |
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:
68833
diff
changeset
|
301 |
using insert by auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
302 |
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:
68833
diff
changeset
|
303 |
show ?case |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
304 |
proof (intro exI conjI) |
69745 | 305 |
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:
68833
diff
changeset
|
306 |
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:
68833
diff
changeset
|
307 |
next |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
308 |
show "finite {i \<in> I. ?U i \<noteq> topspace (X i)}" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
309 |
by (rule rev_finite_subset [OF finite.insertI [OF fin]]) auto |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
310 |
next |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
311 |
show "\<forall>i\<in>I. openin (X i) (?U i)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
312 |
by (simp add: \<open>openin (X i) V\<close> ope openin_Int) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
313 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
314 |
qed (auto intro: dest: not_finite_existsD) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
315 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
316 |
moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
317 |
proof clarify |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
318 |
fix U :: "'a \<Rightarrow> 'b set" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
319 |
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:
68833
diff
changeset
|
320 |
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:
68833
diff
changeset
|
321 |
show "?lhs (Pi\<^sub>E I U)" |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
322 |
unfolding relative_to_def topspace_product_topology |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
323 |
proof (intro exI conjI) |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
324 |
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:
68833
diff
changeset
|
325 |
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:
68833
diff
changeset
|
326 |
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:
68833
diff
changeset
|
327 |
using ope openin_subset by fastforce |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
328 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
329 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
330 |
ultimately show ?thesis |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
331 |
by meson |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
332 |
qed |
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
333 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
334 |
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:
69922
diff
changeset
|
335 |
"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:
69922
diff
changeset
|
336 |
(\<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:
69922
diff
changeset
|
337 |
(\<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:
69922
diff
changeset
|
338 |
unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology |
76838
04c7ec38874e
removed an unfortunate sledgehammer command
paulson <lp15@cam.ac.uk>
parents:
76821
diff
changeset
|
339 |
by (smt (verit, best)) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
340 |
|
69874 | 341 |
lemma closure_of_product_topology: |
342 |
"(product_topology X I) closure_of (PiE I S) = PiE I (\<lambda>i. (X i) closure_of (S i))" |
|
343 |
proof - |
|
344 |
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)) |
|
345 |
\<longleftrightarrow> (\<forall>i \<in> I. \<forall>T. f i \<in> T \<and> openin (X i) T \<longrightarrow> S i \<inter> T \<noteq> {})" |
|
346 |
(is "?lhs = ?rhs") |
|
347 |
if top: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i)" and ext: "f \<in> extensional I" for f |
|
348 |
proof |
|
349 |
assume L[rule_format]: ?lhs |
|
350 |
show ?rhs |
|
351 |
proof clarify |
|
352 |
fix i T |
|
353 |
assume "i \<in> I" "f i \<in> T" "openin (X i) T" "S i \<inter> T = {}" |
|
354 |
then have "openin (product_topology X I) ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> {x. x i \<in> T})" |
|
355 |
by (force simp: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc) |
|
356 |
then show "False" |
|
357 |
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> |
|
358 |
by (auto simp: top ext PiE_iff) |
|
359 |
qed |
|
360 |
next |
|
361 |
assume R [rule_format]: ?rhs |
|
362 |
show ?lhs |
|
363 |
proof (clarsimp simp: openin_product_topology union_of_def arbitrary_def) |
|
364 |
fix \<U> U |
|
365 |
assume |
|
366 |
\<U>: "\<U> \<subseteq> Collect |
|
367 |
(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 |
|
368 |
(\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and |
|
369 |
"f \<in> U" "U \<in> \<U>" |
|
370 |
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) |
|
371 |
relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))) U" |
|
372 |
by blast |
|
373 |
with \<open>f \<in> U\<close> \<open>U \<in> \<U>\<close> |
|
374 |
obtain \<T> where "finite \<T>" |
|
375 |
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}" |
|
376 |
and "topspace (product_topology X I) \<inter> \<Inter> \<T> \<subseteq> U" "f \<in> topspace (product_topology X I) \<inter> \<Inter> \<T>" |
|
377 |
apply (clarsimp simp add: relative_to_def intersection_of_def) |
|
378 |
apply (rule that, auto dest!: subsetD) |
|
379 |
done |
|
380 |
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" |
|
381 |
by (auto simp: PiE_iff) |
|
382 |
have *: "f i \<in> topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>} |
|
383 |
\<and> openin (X i) (topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>})" |
|
384 |
if "i \<in> I" for i |
|
385 |
proof - |
|
386 |
have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)" |
|
387 |
proof (rule finite_vimageI [OF \<open>finite \<T>\<close>]) |
|
388 |
show "inj (\<lambda>U. {x. x i \<in> U})" |
|
389 |
by (auto simp: inj_on_def) |
|
390 |
qed |
|
391 |
then have fin: "finite {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}" |
|
392 |
by (rule rev_finite_subset) auto |
|
393 |
have "openin (X i) (\<Inter> (insert (topspace (X i)) {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}))" |
|
394 |
by (rule openin_Inter) (auto simp: fin) |
|
395 |
then show ?thesis |
|
396 |
using \<open>f \<in> \<Inter> \<T>\<close> by (fastforce simp: that top) |
|
397 |
qed |
|
398 |
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>}" |
|
399 |
have "\<forall>i \<in> I. \<exists>x. x \<in> S i \<inter> \<Phi> i" |
|
400 |
using R [OF _ *] unfolding \<Phi>_def by blast |
|
401 |
then obtain \<theta> where \<theta> [rule_format]: "\<forall>i \<in> I. \<theta> i \<in> S i \<inter> \<Phi> i" |
|
402 |
by metis |
|
403 |
show "\<exists>y\<in>Pi\<^sub>E I S. \<exists>x\<in>\<U>. y \<in> x" |
|
404 |
proof |
|
405 |
show "\<exists>U \<in> \<U>. (\<lambda>i \<in> I. \<theta> i) \<in> U" |
|
406 |
proof |
|
407 |
have "restrict \<theta> I \<in> PiE I (topspace \<circ> X) \<inter> \<Inter>\<T>" |
|
408 |
using \<T> by (fastforce simp: \<Phi>_def PiE_def dest: \<theta>) |
|
409 |
then show "restrict \<theta> I \<in> U" |
|
410 |
using subU by blast |
|
411 |
qed (rule \<open>U \<in> \<U>\<close>) |
|
412 |
next |
|
413 |
show "(\<lambda>i \<in> I. \<theta> i) \<in> Pi\<^sub>E I S" |
|
414 |
using \<theta> by simp |
|
415 |
qed |
|
416 |
qed |
|
417 |
qed |
|
418 |
show ?thesis |
|
419 |
apply (simp add: * closure_of_def PiE_iff set_eq_iff cong: conj_cong) |
|
420 |
by metis |
|
421 |
qed |
|
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
422 |
|
69874 | 423 |
corollary closedin_product_topology: |
424 |
"closedin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. closedin (X i) (S i))" |
|
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
425 |
by (smt (verit, best) PiE_eq closedin_empty closure_of_eq closure_of_product_topology) |
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
426 |
|
69874 | 427 |
corollary closedin_product_topology_singleton: |
428 |
"f \<in> extensional I \<Longrightarrow> closedin (product_topology X I) {f} \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) {f i})" |
|
429 |
using PiE_singleton closedin_product_topology [of X I] |
|
430 |
by (metis (no_types, lifting) all_not_in_conv insertI1) |
|
69030
1eb517214deb
more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents:
68833
diff
changeset
|
431 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
432 |
lemma product_topology_empty: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
433 |
"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:
69918
diff
changeset
|
434 |
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:
69918
diff
changeset
|
435 |
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:
69918
diff
changeset
|
436 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
437 |
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:
69918
diff
changeset
|
438 |
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:
69918
diff
changeset
|
439 |
by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
440 |
|
69874 | 441 |
|
442 |
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
|
443 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
444 |
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
|
445 |
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:
69922
diff
changeset
|
446 |
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
|
447 |
proof - |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
448 |
{ |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
449 |
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
|
450 |
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
|
451 |
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 | 452 |
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
|
453 |
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
|
454 |
have **: "(\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}" |
64911 | 455 |
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
|
456 |
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
|
457 |
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
|
458 |
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
|
459 |
apply (subst *) |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
460 |
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
|
461 |
} |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
462 |
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:
69922
diff
changeset
|
463 |
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
|
464 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
465 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
466 |
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:
69922
diff
changeset
|
467 |
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
|
468 |
"\<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:
69922
diff
changeset
|
469 |
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
|
470 |
unfolding product_topology_def |
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
471 |
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
|
472 |
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
|
473 |
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
|
474 |
by blast |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
475 |
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
|
476 |
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
|
477 |
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:
69922
diff
changeset
|
478 |
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
|
479 |
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
|
480 |
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
|
481 |
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
|
482 |
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
|
483 |
also have "... = (\<Inter>i\<in>J. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)" |
64911 | 484 |
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
|
485 |
also have "openin T1 (...)" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
486 |
using H(2) \<open>J \<subseteq> I\<close> \<open>finite J\<close> assms(1) by blast |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
487 |
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
|
488 |
next |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
489 |
have "f ` topspace T1 \<subseteq> topspace (product_topology T I)" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
490 |
using assms continuous_map_def by fastforce |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
491 |
then 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)}}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
492 |
by (simp add: product_topology_def) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
493 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
494 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
495 |
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:
69922
diff
changeset
|
496 |
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:
69922
diff
changeset
|
497 |
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
|
498 |
"\<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
|
499 |
proof - |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
500 |
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
|
501 |
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:
69922
diff
changeset
|
502 |
also have "continuous_map T1 (T i) (...)" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
503 |
by (metis \<open>i \<in> I\<close> assms continuous_map_compose continuous_map_product_coordinates) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
504 |
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
|
505 |
by simp |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
506 |
next |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
507 |
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
|
508 |
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:
69922
diff
changeset
|
509 |
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
|
510 |
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:
68833
diff
changeset
|
511 |
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
|
512 |
then show "f x i = undefined" |
64911 | 513 |
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
|
514 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
515 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
516 |
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
|
517 |
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:
69922
diff
changeset
|
518 |
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:
69922
diff
changeset
|
519 |
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
|
520 |
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
|
521 |
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:
69922
diff
changeset
|
522 |
then show "continuous_map (product_topology T I) (T i) (\<lambda>x. restrict x J i)" |
64911 | 523 |
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
|
524 |
next |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
525 |
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
|
526 |
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
|
527 |
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
|
528 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
529 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
530 |
|
70136 | 531 |
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
|
532 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
533 |
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
|
534 |
begin |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
535 |
|
70136 | 536 |
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
|
537 |
"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
|
538 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
539 |
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
|
540 |
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:
68833
diff
changeset
|
541 |
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
|
542 |
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
|
543 |
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
|
544 |
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
|
545 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
546 |
end |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
547 |
|
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
548 |
lemma open_PiE [intro?]: |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
549 |
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:
70019
diff
changeset
|
550 |
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:
70019
diff
changeset
|
551 |
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:
70019
diff
changeset
|
552 |
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:
70019
diff
changeset
|
553 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
554 |
lemma euclidean_product_topology: |
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
555 |
"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
|
556 |
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
|
557 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
558 |
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
|
559 |
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
|
560 |
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
|
561 |
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:
69680
diff
changeset
|
562 |
proof - |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
563 |
define V where "V \<equiv> (\<lambda>y. if y \<in> x`I then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
564 |
define X where "X \<equiv> (\<lambda>y. if y \<in> x`I then V y else UNIV)" |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
565 |
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
|
566 |
unfolding X_def V_def using assms by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
567 |
then have "open (Pi\<^sub>E UNIV X)" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
568 |
by (simp add: X_def assms(1) 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
|
569 |
moreover have "Pi\<^sub>E UNIV X = {f. \<forall>i\<in>I. f (x i) \<in> U i}" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
570 |
by (fastforce simp add: PiE_iff X_def V_def split: if_split_asm) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
571 |
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
|
572 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
573 |
|
64911 | 574 |
text \<open>The results proved in the general situation of products of possibly different |
575 |
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
|
576 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
577 |
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
|
578 |
"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:
69922
diff
changeset
|
579 |
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:
69922
diff
changeset
|
580 |
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
|
581 |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
582 |
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:
69918
diff
changeset
|
583 |
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:
69874
diff
changeset
|
584 |
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:
69874
diff
changeset
|
585 |
shows "continuous_on S f" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
586 |
by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
587 |
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
|
588 |
|
69918
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
589 |
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:
69874
diff
changeset
|
590 |
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:
69874
diff
changeset
|
591 |
shows "continuous_on S (\<lambda>x. f x i)" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
592 |
by (metis UNIV_I assms continuous_map_iff_continuous continuous_map_product_then_coordinatewise(1) euclidean_product_topology) |
69918
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
593 |
|
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
594 |
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:
69874
diff
changeset
|
595 |
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:
69874
diff
changeset
|
596 |
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:
69994
diff
changeset
|
597 |
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
|
598 |
|
70086
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
70065
diff
changeset
|
599 |
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:
71172
diff
changeset
|
600 |
fixes B :: "'a::real_normed_vector set" |
70086
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
70065
diff
changeset
|
601 |
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:
70065
diff
changeset
|
602 |
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:
70065
diff
changeset
|
603 |
proof (rule continuous_map_euclidean_top_of_set) |
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
70065
diff
changeset
|
604 |
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:
70065
diff
changeset
|
605 |
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:
70065
diff
changeset
|
606 |
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:
70065
diff
changeset
|
607 |
by (intro continuous_intros) auto |
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
70065
diff
changeset
|
608 |
qed |
72c52a897de2
First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
70065
diff
changeset
|
609 |
|
70136 | 610 |
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
|
611 |
|
64911 | 612 |
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
|
613 |
of product spaces, but they have more to do with countability and could |
64911 | 614 |
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
|
615 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
616 |
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
|
617 |
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
|
618 |
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
|
619 |
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
|
620 |
proof - |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
621 |
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
|
622 |
\<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
|
623 |
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
|
624 |
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
|
625 |
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
|
626 |
case 0 |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
627 |
have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>(0::nat). x i = a)} = {(\<lambda>i. a)}" |
64911 | 628 |
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
|
629 |
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
|
630 |
next |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
631 |
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
|
632 |
define f::"((nat \<Rightarrow> 'a) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
633 |
where "f = (\<lambda>(x, b). x(N:=b))" |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
634 |
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
|
635 |
proof (auto) |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
636 |
fix x assume H: "\<forall>i::nat. x i \<in> F" "\<forall>i\<ge>Suc N. x i = a" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
637 |
have "f (x(N:=a), x N) = x" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
638 |
unfolding f_def by auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
639 |
moreover have "(x(N:=a), x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
640 |
using H \<open>a \<in> F\<close> by auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
641 |
ultimately show "x \<in> f ` ({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)" |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
642 |
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
|
643 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
644 |
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
|
645 |
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
|
646 |
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
|
647 |
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
|
648 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
649 |
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
|
650 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
651 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
652 |
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
|
653 |
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
|
654 |
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
|
655 |
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
|
656 |
proof - |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
657 |
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
|
658 |
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
|
659 |
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
|
660 |
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
|
661 |
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
|
662 |
\<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
|
663 |
proof (auto) |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
664 |
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
|
665 |
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
|
666 |
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
|
667 |
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
|
668 |
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
|
669 |
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
|
670 |
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
|
671 |
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
|
672 |
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
|
673 |
by auto |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
674 |
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
|
675 |
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
|
676 |
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
|
677 |
by auto |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
678 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
679 |
then show ?thesis |
64911 | 680 |
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
|
681 |
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
|
682 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
683 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
684 |
instance "fun" :: (countable, first_countable_topology) first_countable_topology |
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
685 |
proof |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
686 |
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
|
687 |
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
|
688 |
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
|
689 |
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
|
690 |
"\<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
|
691 |
"\<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
|
692 |
by metis |
69566 | 693 |
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
|
694 |
define B where "B = (\<lambda>i. (A (x i))`UNIV \<union> {UNIV})" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
695 |
have countB: "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:
70019
diff
changeset
|
696 |
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:
70019
diff
changeset
|
697 |
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
|
698 |
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
|
699 |
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
|
700 |
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
|
701 |
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
|
702 |
have "countable {X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
703 |
by (simp add: countB 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
|
704 |
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
|
705 |
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
|
706 |
ultimately have "countable K" by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
707 |
have I: "x \<in> k" if "k \<in> K" for k |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
708 |
using that unfolding K_def B_def apply auto using A(1) by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
709 |
have II: "open k" if "k \<in> K" for k |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
710 |
using that unfolding K_def by (blast intro: open_B 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
|
711 |
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
|
712 |
proof - |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
713 |
have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U" |
64911 | 714 |
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
|
715 |
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
|
716 |
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:
69683
diff
changeset
|
717 |
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
|
718 |
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
|
719 |
"\<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
|
720 |
"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
|
721 |
"(\<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
|
722 |
by auto |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
723 |
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
|
724 |
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
|
725 |
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
|
726 |
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
|
727 |
have **: "Y i \<in> B i \<and> Y i \<subseteq> X i" for i |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
728 |
proof (cases "i \<in> I") |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
729 |
case True |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
730 |
then show ?thesis |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
731 |
by (metis (mono_tags, lifting) "*" Nitpick.Eps_psimp Y_def) |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
732 |
next |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
733 |
case False |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
734 |
then show ?thesis by (simp add: B_def I_def Y_def) |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
735 |
qed |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
736 |
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
|
737 |
unfolding Y_def by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
738 |
with ** have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
739 |
using H(3) I_def finite_subset by blast |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
740 |
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
|
741 |
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
|
742 |
have "Y i \<subseteq> X i" for i |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
743 |
using "**" by auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
744 |
then have "Pi\<^sub>E UNIV Y \<subseteq> U" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
745 |
by (metis H(4) PiE_mono subset_trans) |
64911 | 746 |
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
|
747 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
748 |
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))" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
749 |
using \<open>countable K\<close> I II Inc by (simp add: first_countableI) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
750 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
751 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
752 |
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
|
753 |
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
|
754 |
topological_basis K \<and> countable K \<and> |
64910 | 755 |
(\<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:
69680
diff
changeset
|
756 |
proof - |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
757 |
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
|
758 |
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
|
759 |
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
|
760 |
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
|
761 |
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
|
762 |
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
|
763 |
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
|
764 |
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
|
765 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
766 |
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 | 767 |
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 | 768 |
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
|
769 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
770 |
have "countable {X. (\<forall>(i::'a). X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
771 |
using \<open>countable B2\<close> by (intro countable_product_event_const) auto |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
772 |
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
|
773 |
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
|
774 |
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
|
775 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
776 |
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
|
777 |
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
|
778 |
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
|
779 |
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
|
780 |
unfolding open_fun_def by auto |
64911 | 781 |
with product_topology_open_contains_basis[OF this \<open>x \<in> U\<close>] |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
782 |
obtain X where H: "x \<in> (\<Pi>\<^sub>E i\<in>UNIV. 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
|
783 |
"\<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
|
784 |
"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
|
785 |
"(\<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
|
786 |
by auto |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
787 |
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
|
788 |
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
|
789 |
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
|
790 |
have *: "\<exists>y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y" for i |
64911 | 791 |
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
|
792 |
have **: "Y i \<in> B2 \<and> Y i \<subseteq> X i \<and> x i \<in> Y i" for i |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
793 |
using someI_ex[OF *] by (simp add: B2_def I_def Y_def) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
794 |
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
|
795 |
unfolding Y_def by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
796 |
then have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
797 |
using "**" H(3) I_def finite_subset by blast |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
798 |
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
|
799 |
unfolding K_def by auto |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
800 |
then show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
801 |
by (meson "**" H(4) PiE_I PiE_mono UNIV_I order.trans) |
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
802 |
next |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
803 |
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
|
804 |
show "open U" |
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
805 |
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
|
806 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
807 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
808 |
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
|
809 |
qed |
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
810 |
|
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
811 |
instance "fun" :: (countable, second_countable_topology) second_countable_topology |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
812 |
proof |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
813 |
show "\<exists>B::('a \<Rightarrow> 'b) set set. countable B \<and> open = generate_topology B" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
814 |
using product_topology_countable_basis topological_basis_imp_subbasis |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
815 |
by auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
816 |
qed |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
817 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
818 |
subsection\<open>The Alexander subbase theorem\<close> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
819 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
820 |
theorem Alexander_subbase: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
821 |
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:
69918
diff
changeset
|
822 |
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:
69918
diff
changeset
|
823 |
shows "compact_space X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
824 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
825 |
have U\<B>: "\<Union>\<B> = topspace X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
826 |
by (simp flip: X) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
827 |
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:
69918
diff
changeset
|
828 |
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:
69918
diff
changeset
|
829 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
830 |
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:
69918
diff
changeset
|
831 |
have 1: "\<A> \<noteq> {}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
832 |
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:
69918
diff
changeset
|
833 |
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:
69918
diff
changeset
|
834 |
unfolding \<A>_def |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
835 |
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:
69918
diff
changeset
|
836 |
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:
69918
diff
changeset
|
837 |
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:
69918
diff
changeset
|
838 |
have "\<C> \<subseteq> \<A>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
839 |
using subset_chain_def \<C> by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
840 |
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:
69918
diff
changeset
|
841 |
by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
842 |
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:
69918
diff
changeset
|
843 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
844 |
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:
69918
diff
changeset
|
845 |
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:
69918
diff
changeset
|
846 |
then show False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
847 |
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:
69918
diff
changeset
|
848 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
849 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
850 |
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:
69918
diff
changeset
|
851 |
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:
69918
diff
changeset
|
852 |
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:
69918
diff
changeset
|
853 |
\<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:
69918
diff
changeset
|
854 |
\<Longrightarrow> \<W> = \<K>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
855 |
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:
69918
diff
changeset
|
856 |
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:
69918
diff
changeset
|
857 |
unfolding \<A>_def by simp_all metis+ |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
858 |
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:
69918
diff
changeset
|
859 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
860 |
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:
69918
diff
changeset
|
861 |
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:
69918
diff
changeset
|
862 |
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:
69918
diff
changeset
|
863 |
by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
864 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
865 |
using that by (metis U\<B>) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
866 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
867 |
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:
69918
diff
changeset
|
868 |
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:
69918
diff
changeset
|
869 |
then have "C \<subseteq> topspace X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
870 |
by (metis openin_subset) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
871 |
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:
69918
diff
changeset
|
872 |
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:
69918
diff
changeset
|
873 |
moreover have "C \<noteq> topspace X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
874 |
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:
69918
diff
changeset
|
875 |
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:
69918
diff
changeset
|
876 |
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:
69918
diff
changeset
|
877 |
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:
69918
diff
changeset
|
878 |
then have "\<Union>\<V> \<subseteq> topspace X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
879 |
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:
69918
diff
changeset
|
880 |
then have "topspace X \<notin> \<V>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
881 |
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:
69918
diff
changeset
|
882 |
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:
69918
diff
changeset
|
883 |
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:
69918
diff
changeset
|
884 |
then have "\<Inter>\<B>' \<subseteq> \<Union>\<B>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
885 |
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:
69918
diff
changeset
|
886 |
then have "\<Inter>\<B>' \<subseteq> C" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
887 |
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:
69918
diff
changeset
|
888 |
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:
69918
diff
changeset
|
889 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
890 |
fix b |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
891 |
assume "b \<in> \<B>'" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
892 |
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:
69918
diff
changeset
|
893 |
proof (rule *) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
894 |
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:
69918
diff
changeset
|
895 |
using that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
896 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
897 |
have "b \<in> \<B>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
898 |
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:
69918
diff
changeset
|
899 |
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:
69918
diff
changeset
|
900 |
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:
69918
diff
changeset
|
901 |
moreover have "\<Union>\<B> \<inter> b = b" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
902 |
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:
69918
diff
changeset
|
903 |
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:
69918
diff
changeset
|
904 |
using that \<open>b \<in> \<B>'\<close> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
905 |
apply (simp add: openin_subbase flip: X) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
906 |
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:
69918
diff
changeset
|
907 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
908 |
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:
69918
diff
changeset
|
909 |
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:
69918
diff
changeset
|
910 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
911 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
912 |
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:
69918
diff
changeset
|
913 |
using top by auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
914 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
915 |
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:
69918
diff
changeset
|
916 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
917 |
have "insert b (\<F> \<inter> \<K>) = \<F>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
918 |
using non that by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
919 |
then show False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
920 |
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:
69918
diff
changeset
|
921 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
922 |
qed auto |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
923 |
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:
69918
diff
changeset
|
924 |
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:
69918
diff
changeset
|
925 |
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:
69918
diff
changeset
|
926 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
927 |
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:
69918
diff
changeset
|
928 |
by metis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
929 |
let ?\<D> = "insert C (\<Union>(F ` \<B>'))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
930 |
show False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
931 |
proof (rule non) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
932 |
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:
69918
diff
changeset
|
933 |
using F by (simp add: INT_greatest) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
934 |
also have "\<dots> \<subseteq> \<Union>?\<D>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
935 |
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:
69918
diff
changeset
|
936 |
finally show "topspace X \<subseteq> \<Union>?\<D>" . |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
937 |
show "?\<D> \<subseteq> \<K>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
938 |
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:
69918
diff
changeset
|
939 |
show "finite ?\<D>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
940 |
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:
69918
diff
changeset
|
941 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
942 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
943 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
944 |
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:
69918
diff
changeset
|
945 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
946 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
947 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
948 |
corollary Alexander_subbase_alt: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
949 |
assumes "U \<subseteq> \<Union>\<B>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
950 |
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:
69918
diff
changeset
|
951 |
and X: "topology |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
952 |
(arbitrary union_of |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
953 |
(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:
69918
diff
changeset
|
954 |
shows "compact_space X" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
955 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
956 |
have "topspace X = U" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
957 |
using X topspace_subbase by fastforce |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
958 |
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:
69918
diff
changeset
|
959 |
unfolding relative_to_def |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
960 |
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:
69918
diff
changeset
|
961 |
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:
69918
diff
changeset
|
962 |
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:
69918
diff
changeset
|
963 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
964 |
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:
69918
diff
changeset
|
965 |
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:
69918
diff
changeset
|
966 |
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:
69918
diff
changeset
|
967 |
by (auto simp: subset_image_iff) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
968 |
moreover have "U \<subseteq> \<Union>\<B>'" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
969 |
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:
69918
diff
changeset
|
970 |
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:
69918
diff
changeset
|
971 |
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:
69918
diff
changeset
|
972 |
then show ?thesis |
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
973 |
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:
69918
diff
changeset
|
974 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
975 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
976 |
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:
69918
diff
changeset
|
977 |
apply (simp flip: X) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
978 |
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:
69918
diff
changeset
|
979 |
apply (blast intro: *) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
980 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
981 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
982 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
983 |
proposition continuous_map_componentwise: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
984 |
"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:
69918
diff
changeset
|
985 |
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:
69918
diff
changeset
|
986 |
(is "?lhs \<longleftrightarrow> _ \<and> ?rhs") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
987 |
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:
69918
diff
changeset
|
988 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
989 |
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:
69918
diff
changeset
|
990 |
by force |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
991 |
moreover have ?rhs if L: ?lhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
992 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
993 |
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:
69918
diff
changeset
|
994 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
995 |
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:
69918
diff
changeset
|
996 |
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:
69918
diff
changeset
|
997 |
apply (simp add: relative_to_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
998 |
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:
69918
diff
changeset
|
999 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1000 |
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:
69918
diff
changeset
|
1001 |
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:
69918
diff
changeset
|
1002 |
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:
69918
diff
changeset
|
1003 |
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:
69918
diff
changeset
|
1004 |
ultimately show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1005 |
by metis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1006 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1007 |
with that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1008 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1009 |
by (auto simp: continuous_map_def) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1010 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1011 |
moreover have ?lhs if ?rhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1012 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1013 |
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:
69918
diff
changeset
|
1014 |
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:
69918
diff
changeset
|
1015 |
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:
69918
diff
changeset
|
1016 |
by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1017 |
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:
69918
diff
changeset
|
1018 |
by blast |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1019 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1020 |
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:
69918
diff
changeset
|
1021 |
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:
69918
diff
changeset
|
1022 |
fix \<T> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1023 |
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:
69918
diff
changeset
|
1024 |
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:
69918
diff
changeset
|
1025 |
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:
69918
diff
changeset
|
1026 |
proof (rule openin_Union; clarify) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1027 |
fix S T |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1028 |
assume "T \<in> \<T>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1029 |
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:
69918
diff
changeset
|
1030 |
"\<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:
69918
diff
changeset
|
1031 |
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:
69918
diff
changeset
|
1032 |
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:
69918
diff
changeset
|
1033 |
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:
69918
diff
changeset
|
1034 |
unfolding 3 |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1035 |
apply (rule openin_Inter; auto) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1036 |
done |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1037 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1038 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1039 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1040 |
ultimately show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1041 |
by metis |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1042 |
qed (auto simp: continuous_map_def PiE_def) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1043 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1044 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1045 |
lemma continuous_map_componentwise_UNIV: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1046 |
"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:
69918
diff
changeset
|
1047 |
by (simp add: continuous_map_componentwise) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1048 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1049 |
lemma continuous_map_product_projection [continuous_intros]: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1050 |
"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:
69918
diff
changeset
|
1051 |
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:
69918
diff
changeset
|
1052 |
|
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1053 |
declare continuous_map_from_subtopology [OF continuous_map_product_projection, continuous_intros] |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1054 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1055 |
proposition open_map_product_projection: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1056 |
assumes "i \<in> I" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1057 |
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:
69918
diff
changeset
|
1058 |
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:
69918
diff
changeset
|
1059 |
proof clarify |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1060 |
fix \<V> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1061 |
assume \<V>: "\<V> \<subseteq> Collect |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1062 |
(finite intersection_of |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1063 |
(\<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:
69918
diff
changeset
|
1064 |
topspace (product_topology Y I))" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1065 |
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:
69918
diff
changeset
|
1066 |
proof (rule openin_Union, clarify) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1067 |
fix S V |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1068 |
assume "V \<in> \<V>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1069 |
obtain \<F> where "finite \<F>" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1070 |
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:
69918
diff
changeset
|
1071 |
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:
69918
diff
changeset
|
1072 |
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:
69918
diff
changeset
|
1073 |
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:
69918
diff
changeset
|
1074 |
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:
69918
diff
changeset
|
1075 |
proof (subst openin_subopen; clarify) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1076 |
fix x f |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1077 |
assume "f \<in> V" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1078 |
let ?T = "{a \<in> topspace(Y i). |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1079 |
(\<lambda>j\<in>I. f j)(i:=a) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1080 |
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:
69918
diff
changeset
|
1081 |
proof (intro exI conjI) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1082 |
show "openin (Y i) ?T" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1083 |
proof (rule openin_continuous_map_preimage) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1084 |
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:
69918
diff
changeset
|
1085 |
proof (cases "k=i") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1086 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1087 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1088 |
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:
69918
diff
changeset
|
1089 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1090 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1091 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1092 |
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:
69918
diff
changeset
|
1093 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1094 |
then show "continuous_map (Y i) (product_topology Y I) |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1095 |
(\<lambda>x. (\<lambda>j\<in>I. f j)(i:=x))" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1096 |
by (auto simp: continuous_map_componentwise assms extensional_def restrict_def) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1097 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1098 |
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:
69918
diff
changeset
|
1099 |
by (metis openin_topspace topspace_product_topology) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1100 |
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:
69918
diff
changeset
|
1101 |
if "\<F> \<noteq> {}" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1102 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1103 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1104 |
proof (rule openin_Inter) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1105 |
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:
69918
diff
changeset
|
1106 |
unfolding openin_product_topology relative_to_def |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1107 |
apply (clarify intro!: arbitrary_union_of_inc) |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1108 |
using subsetD [OF \<F>] |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1109 |
by (metis (mono_tags, lifting) finite_intersection_of_inc mem_Collect_eq topspace_product_topology) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1110 |
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:
69918
diff
changeset
|
1111 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1112 |
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:
69918
diff
changeset
|
1113 |
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:
69918
diff
changeset
|
1114 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1115 |
next |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1116 |
have eqf: "(\<lambda>j\<in>I. f j)(i:=f i) = f" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1117 |
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:
69918
diff
changeset
|
1118 |
show "f i \<in> ?T" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1119 |
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:
69918
diff
changeset
|
1120 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1121 |
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:
69918
diff
changeset
|
1122 |
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:
69918
diff
changeset
|
1123 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1124 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1125 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1126 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1127 |
|
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1128 |
lemma retraction_map_product_projection: |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1129 |
assumes "i \<in> I" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1130 |
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:
69918
diff
changeset
|
1131 |
(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:
69918
diff
changeset
|
1132 |
(is "?lhs = ?rhs") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1133 |
proof |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1134 |
assume ?lhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1135 |
then show ?rhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1136 |
using retraction_imp_surjective_map by force |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1137 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1138 |
assume R: ?rhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1139 |
show ?lhs |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1140 |
proof (cases "topspace (product_topology X I) = {}") |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1141 |
case True |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1142 |
then show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1143 |
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:
69918
diff
changeset
|
1144 |
next |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1145 |
case False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1146 |
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:
69918
diff
changeset
|
1147 |
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:
69918
diff
changeset
|
1148 |
proof - |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1149 |
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:
69918
diff
changeset
|
1150 |
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:
69918
diff
changeset
|
1151 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1152 |
using \<open>i \<in> I\<close> that |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1153 |
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:
69918
diff
changeset
|
1154 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1155 |
show ?thesis |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1156 |
using \<open>i \<in> I\<close> False |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1157 |
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:
69918
diff
changeset
|
1158 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1159 |
qed |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1160 |
|
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1161 |
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:
69922
diff
changeset
|
1162 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1163 |
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:
69922
diff
changeset
|
1164 |
"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:
69922
diff
changeset
|
1165 |
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:
69922
diff
changeset
|
1166 |
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:
69922
diff
changeset
|
1167 |
(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:
69922
diff
changeset
|
1168 |
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:
69922
diff
changeset
|
1169 |
case False |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1170 |
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:
69922
diff
changeset
|
1171 |
proof |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1172 |
assume L: ?lhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1173 |
moreover |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1174 |
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:
69922
diff
changeset
|
1175 |
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:
69922
diff
changeset
|
1176 |
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:
69922
diff
changeset
|
1177 |
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:
69922
diff
changeset
|
1178 |
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:
69922
diff
changeset
|
1179 |
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:
69922
diff
changeset
|
1180 |
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:
69922
diff
changeset
|
1181 |
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:
69922
diff
changeset
|
1182 |
show ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1183 |
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:
69922
diff
changeset
|
1184 |
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:
69922
diff
changeset
|
1185 |
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:
69922
diff
changeset
|
1186 |
using * |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1187 |
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:
69922
diff
changeset
|
1188 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1189 |
fix i :: "'a" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1190 |
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:
69922
diff
changeset
|
1191 |
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:
69922
diff
changeset
|
1192 |
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:
69922
diff
changeset
|
1193 |
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:
69922
diff
changeset
|
1194 |
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:
69922
diff
changeset
|
1195 |
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:
69922
diff
changeset
|
1196 |
done |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1197 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1198 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1199 |
assume ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1200 |
then show ?lhs |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1201 |
unfolding openin_product_topology |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1202 |
by (intro arbitrary_union_of_inc) (auto simp: product_topology_base_alt) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1203 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1204 |
ultimately show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1205 |
by simp |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1206 |
qed simp |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1207 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1208 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1209 |
corollary openin_PiE: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1210 |
"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:
69922
diff
changeset
|
1211 |
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:
69922
diff
changeset
|
1212 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1213 |
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:
69922
diff
changeset
|
1214 |
"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:
69922
diff
changeset
|
1215 |
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:
69922
diff
changeset
|
1216 |
(is "?lhs = ?rhs") |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1217 |
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:
69922
diff
changeset
|
1218 |
case False |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1219 |
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:
69922
diff
changeset
|
1220 |
by auto |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1221 |
show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1222 |
proof |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1223 |
assume L: ?lhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1224 |
show ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1225 |
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:
69922
diff
changeset
|
1226 |
fix i |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1227 |
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:
69922
diff
changeset
|
1228 |
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:
69922
diff
changeset
|
1229 |
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:
69922
diff
changeset
|
1230 |
moreover |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1231 |
have "\<And>x. x \<in> topspace (X i) \<Longrightarrow> x \<in> (\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i))" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1232 |
using \<open>i \<in> I\<close> z by (rule_tac x="z(i:=x)" in image_eqI) auto |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1233 |
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:
69922
diff
changeset
|
1234 |
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:
69922
diff
changeset
|
1235 |
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:
69922
diff
changeset
|
1236 |
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:
69922
diff
changeset
|
1237 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1238 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1239 |
assume R: ?rhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1240 |
show ?lhs |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1241 |
proof (cases "I = {}") |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1242 |
case True |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1243 |
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:
69922
diff
changeset
|
1244 |
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:
69922
diff
changeset
|
1245 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1246 |
case False |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1247 |
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:
69922
diff
changeset
|
1248 |
by blast |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1249 |
show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1250 |
using R |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1251 |
proof |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1252 |
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:
69922
diff
changeset
|
1253 |
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:
69922
diff
changeset
|
1254 |
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:
69922
diff
changeset
|
1255 |
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:
69922
diff
changeset
|
1256 |
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:
69922
diff
changeset
|
1257 |
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:
69922
diff
changeset
|
1258 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1259 |
fix C |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1260 |
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:
69922
diff
changeset
|
1261 |
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:
69922
diff
changeset
|
1262 |
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:
69922
diff
changeset
|
1263 |
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:
69922
diff
changeset
|
1264 |
case True |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1265 |
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:
69922
diff
changeset
|
1266 |
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:
69922
diff
changeset
|
1267 |
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:
69922
diff
changeset
|
1268 |
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:
69922
diff
changeset
|
1269 |
\<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:
69922
diff
changeset
|
1270 |
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:
69922
diff
changeset
|
1271 |
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:
69922
diff
changeset
|
1272 |
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:
69922
diff
changeset
|
1273 |
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:
69922
diff
changeset
|
1274 |
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:
69922
diff
changeset
|
1275 |
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:
69922
diff
changeset
|
1276 |
unfolding \<D>_def |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1277 |
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:
69922
diff
changeset
|
1278 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1279 |
case False |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1280 |
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:
69922
diff
changeset
|
1281 |
by force |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1282 |
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:
69922
diff
changeset
|
1283 |
by metis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1284 |
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:
69922
diff
changeset
|
1285 |
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:
69922
diff
changeset
|
1286 |
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:
69922
diff
changeset
|
1287 |
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:
69922
diff
changeset
|
1288 |
ultimately show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1289 |
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:
69922
diff
changeset
|
1290 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1291 |
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:
69922
diff
changeset
|
1292 |
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:
69922
diff
changeset
|
1293 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1294 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1295 |
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:
69922
diff
changeset
|
1296 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1297 |
corollary compactin_PiE: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1298 |
"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:
69922
diff
changeset
|
1299 |
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:
69922
diff
changeset
|
1300 |
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:
69922
diff
changeset
|
1301 |
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:
69922
diff
changeset
|
1302 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1303 |
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:
69922
diff
changeset
|
1304 |
"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:
69922
diff
changeset
|
1305 |
\<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:
69922
diff
changeset
|
1306 |
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:
69922
diff
changeset
|
1307 |
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:
69922
diff
changeset
|
1308 |
|
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1309 |
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:
69922
diff
changeset
|
1310 |
"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:
69922
diff
changeset
|
1311 |
unfolding homeomorphic_space |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1312 |
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:
69922
diff
changeset
|
1313 |
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:
69922
diff
changeset
|
1314 |
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:
69922
diff
changeset
|
1315 |
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:
69922
diff
changeset
|
1316 |
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:
69922
diff
changeset
|
1317 |
done |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1318 |
|
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1319 |
subsection\<open>Relationship with connected spaces, paths, etc.\<close> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1320 |
|
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1321 |
proposition connected_space_product_topology: |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1322 |
"connected_space(product_topology X I) \<longleftrightarrow> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1323 |
(\<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:
69939
diff
changeset
|
1324 |
(is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs") |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1325 |
proof (cases ?eq) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1326 |
case False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1327 |
moreover have "?lhs = ?rhs" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1328 |
proof |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1329 |
assume ?lhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1330 |
moreover |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1331 |
have "connectedin(X i) (topspace(X i))" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1332 |
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:
69939
diff
changeset
|
1333 |
proof - |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1334 |
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:
69939
diff
changeset
|
1335 |
by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1336 |
show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1337 |
using connectedin_continuous_map_image [OF cm ci] \<open>i \<in> I\<close> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1338 |
by (simp add: False image_projection_PiE) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1339 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1340 |
ultimately show ?rhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1341 |
by (meson connectedin_topspace) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1342 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1343 |
assume cs [rule_format]: ?rhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1344 |
have False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1345 |
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:
69939
diff
changeset
|
1346 |
and U: "openin (product_topology X I) U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1347 |
and V: "openin (product_topology X I) V" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1348 |
and "U \<noteq> {}" "V \<noteq> {}" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1349 |
for U V |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1350 |
proof - |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1351 |
obtain f where "f \<in> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1352 |
using \<open>U \<noteq> {}\<close> by blast |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1353 |
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:
69939
diff
changeset
|
1354 |
using U openin_subset by fastforce |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1355 |
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:
69939
diff
changeset
|
1356 |
using U V openin_subset by blast+ |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1357 |
moreover have "(\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<subseteq> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1358 |
proof - |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1359 |
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:
69939
diff
changeset
|
1360 |
(\<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:
69939
diff
changeset
|
1361 |
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:
69939
diff
changeset
|
1362 |
then obtain \<T> where "finite \<T>" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1363 |
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:
69939
diff
changeset
|
1364 |
and subU: "topspace (product_topology X I) \<inter> \<Inter>\<T> \<subseteq> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1365 |
and ftop: "f \<in> topspace (product_topology X I)" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1366 |
and fint: "f \<in> \<Inter> \<T>" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1367 |
by (fastforce simp: relative_to_def intersection_of_def subset_iff) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1368 |
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:
69939
diff
changeset
|
1369 |
obtain L where "finite L" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1370 |
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:
69939
diff
changeset
|
1371 |
proof |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1372 |
show "finite ?L" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1373 |
proof (rule finite_Union) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1374 |
fix M |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1375 |
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:
69939
diff
changeset
|
1376 |
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:
69939
diff
changeset
|
1377 |
by blast |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1378 |
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:
69939
diff
changeset
|
1379 |
using t by meson |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1380 |
then have "f j \<in> V" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1381 |
using \<open>C \<in> \<T>\<close> fint by force |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1382 |
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:
69939
diff
changeset
|
1383 |
using that |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1384 |
apply (clarsimp simp add: set_eq_iff) |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1385 |
apply (rule_tac x="f(k:=x)" in image_eqI, auto) |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1386 |
done |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1387 |
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:
69939
diff
changeset
|
1388 |
using Ceq by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1389 |
then show "finite M" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1390 |
using C finite_subset by fastforce |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1391 |
qed (use \<open>finite \<T>\<close> in blast) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1392 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1393 |
fix i U |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1394 |
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:
69939
diff
changeset
|
1395 |
then show "i \<in> ?L" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1396 |
by (rule_tac a="{x. x i \<in> U}" in UN_I) (force+) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1397 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1398 |
show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1399 |
proof |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1400 |
fix h |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1401 |
assume h: "h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1402 |
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:
69939
diff
changeset
|
1403 |
have gin: "g \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1404 |
unfolding g_def using f h by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1405 |
moreover have "g \<in> X" if "X \<in> \<T>" for X |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1406 |
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:
69939
diff
changeset
|
1407 |
ultimately have "g \<in> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1408 |
using subU by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1409 |
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:
69939
diff
changeset
|
1410 |
using that |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1411 |
proof (induction arbitrary: h) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1412 |
case empty |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1413 |
then show ?case |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1414 |
using PiE_ext \<open>g \<in> U\<close> gin by force |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1415 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1416 |
case (insert i M) |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1417 |
define f where "f \<equiv> h(i:=g i)" |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1418 |
have fin: "f \<in> PiE I (topspace \<circ> X)" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1419 |
unfolding f_def using gin insert.prems(1) by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1420 |
have subM: "{j \<in> I. f j \<noteq> g j} \<subseteq> M" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1421 |
unfolding f_def using insert.prems(2) by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1422 |
have "f \<in> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1423 |
using insert.IH [OF fin subM] . |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1424 |
show ?case |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1425 |
proof (cases "h \<in> V") |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1426 |
case True |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1427 |
show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1428 |
proof (cases "i \<in> I") |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1429 |
case True |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1430 |
let ?U = "{x \<in> topspace(X i). h(i:=x) \<in> U}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1431 |
let ?V = "{x \<in> topspace(X i). h(i:=x) \<in> V}" |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1432 |
have False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1433 |
proof (rule connected_spaceD [OF cs [OF \<open>i \<in> I\<close>]]) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1434 |
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:
69939
diff
changeset
|
1435 |
using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1436 |
then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x. h(i:=x))" |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1437 |
using \<open>i \<in> I\<close> insert.prems(1) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1438 |
by (auto simp: continuous_map_componentwise extensional_def) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1439 |
show "openin (X i) ?U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1440 |
by (rule openin_continuous_map_preimage [OF cm U]) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1441 |
show "openin (X i) ?V" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1442 |
by (rule openin_continuous_map_preimage [OF cm V]) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1443 |
show "topspace (X i) \<subseteq> ?U \<union> ?V" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1444 |
proof clarsimp |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1445 |
fix x |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1446 |
assume "x \<in> topspace (X i)" and "h(i:=x) \<notin> V" |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1447 |
with True subUV \<open>h \<in> Pi\<^sub>E I (topspace \<circ> X)\<close> |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1448 |
show "h(i:=x) \<in> U" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1449 |
by (force dest: subsetD [where c="h(i:=x)"]) |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1450 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1451 |
show "?U \<inter> ?V = {}" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1452 |
using disj by blast |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1453 |
show "?U \<noteq> {}" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1454 |
using True \<open>f \<in> U\<close> f_def gin by auto |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1455 |
show "?V \<noteq> {}" |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1456 |
using True \<open>h \<in> V\<close> V openin_subset by fastforce |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1457 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1458 |
then show ?thesis .. |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1459 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1460 |
case False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1461 |
show ?thesis |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1462 |
using insert.prems(1) by (metis False gin PiE_E \<open>f \<in> U\<close> f_def fun_upd_triv) |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1463 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1464 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1465 |
case False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1466 |
then show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1467 |
using subUV insert.prems(1) by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1468 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1469 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1470 |
then show "h \<in> U" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1471 |
unfolding g_def using PiE_iff \<open>finite L\<close> h by fastforce |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1472 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1473 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1474 |
ultimately show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1475 |
using disj inf_absorb2 \<open>V \<noteq> {}\<close> by fastforce |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1476 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1477 |
then show ?lhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1478 |
unfolding connected_space_def |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1479 |
by auto |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1480 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1481 |
ultimately show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1482 |
by simp |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1483 |
qed (simp add: connected_space_topspace_empty) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1484 |
|
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1485 |
|
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1486 |
lemma connectedin_PiE: |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1487 |
"connectedin (product_topology X I) (PiE I S) \<longleftrightarrow> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1488 |
PiE I S = {} \<or> (\<forall>i \<in> I. connectedin (X i) (S i))" |
71172 | 1489 |
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:
69939
diff
changeset
|
1490 |
|
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1491 |
lemma path_connected_space_product_topology: |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1492 |
"path_connected_space(product_topology X I) \<longleftrightarrow> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1493 |
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:
69939
diff
changeset
|
1494 |
(is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs") |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1495 |
proof (cases ?eq) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1496 |
case False |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1497 |
moreover have "?lhs = ?rhs" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1498 |
proof |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1499 |
assume L: ?lhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1500 |
show ?rhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1501 |
proof (clarsimp simp flip: path_connectedin_topspace) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1502 |
fix i :: "'a" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1503 |
assume "i \<in> I" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1504 |
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:
69939
diff
changeset
|
1505 |
by (simp add: \<open>i \<in> I\<close> continuous_map_product_projection) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1506 |
show "path_connectedin (X i) (topspace (X i))" |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1507 |
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:
69939
diff
changeset
|
1508 |
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:
69939
diff
changeset
|
1509 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1510 |
next |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1511 |
assume R [rule_format]: ?rhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1512 |
show ?lhs |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1513 |
unfolding path_connected_space_def topspace_product_topology |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1514 |
proof clarify |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1515 |
fix x y |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1516 |
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:
69939
diff
changeset
|
1517 |
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:
69939
diff
changeset
|
1518 |
using PiE_mem R path_connected_space_def x y by force |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1519 |
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:
69939
diff
changeset
|
1520 |
by metis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1521 |
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:
69939
diff
changeset
|
1522 |
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:
69939
diff
changeset
|
1523 |
apply (force simp: pathin_def continuous_map_componentwise) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1524 |
done |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1525 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1526 |
qed |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1527 |
ultimately show ?thesis |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1528 |
by simp |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1529 |
qed (simp add: path_connected_space_topspace_empty) |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1530 |
|
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1531 |
lemma path_connectedin_PiE: |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1532 |
"path_connectedin (product_topology X I) (PiE I S) \<longleftrightarrow> |
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1533 |
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:
69939
diff
changeset
|
1534 |
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:
69939
diff
changeset
|
1535 |
|
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1536 |
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:
69945
diff
changeset
|
1537 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1538 |
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:
69945
diff
changeset
|
1539 |
assumes "i \<in> I" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1540 |
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:
69945
diff
changeset
|
1541 |
(topspace(product_topology X I) = {} \<longrightarrow> topspace(X i) = {})" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1542 |
by (metis assms image_empty quotient_imp_surjective_map retraction_imp_quotient_map |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1543 |
retraction_map_product_projection) |
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1544 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1545 |
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:
69945
diff
changeset
|
1546 |
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:
69945
diff
changeset
|
1547 |
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:
69945
diff
changeset
|
1548 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1549 |
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:
69945
diff
changeset
|
1550 |
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:
69945
diff
changeset
|
1551 |
moreover |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1552 |
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:
69945
diff
changeset
|
1553 |
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:
69945
diff
changeset
|
1554 |
ultimately show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1555 |
unfolding homeomorphic_space_def |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1556 |
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:
69945
diff
changeset
|
1557 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1558 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1559 |
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:
69945
diff
changeset
|
1560 |
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:
69945
diff
changeset
|
1561 |
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:
69945
diff
changeset
|
1562 |
\<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:
69945
diff
changeset
|
1563 |
(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:
69945
diff
changeset
|
1564 |
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:
69945
diff
changeset
|
1565 |
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:
69945
diff
changeset
|
1566 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1567 |
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:
69945
diff
changeset
|
1568 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1569 |
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:
69945
diff
changeset
|
1570 |
by force |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1571 |
have "?SX f i homeomorphic_space X i" |
76821
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1572 |
using f 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})"] |
337c2265d8a2
Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents:
71633
diff
changeset
|
1573 |
by (force simp add: subtopology_PiE) |
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1574 |
then show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1575 |
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:
69945
diff
changeset
|
1576 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69945
diff
changeset
|
1577 |
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:
69945
diff
changeset
|
1578 |
qed |
69945
35ba13ac6e5c
New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
1579 |
|
64289
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff
changeset
|
1580 |
end |