src/HOL/Analysis/Function_Topology.thy
author wenzelm
Wed, 03 Jul 2024 13:54:48 +0200
changeset 80489 e0568c7b5bed
parent 78336 6bae28577994
permissions -rw-r--r--
unused;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69566
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    49
    ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    50
            \<and> (f`(topspace T1) \<subseteq> (topspace T2)))\<close>}
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    51
be the natural continuity definition of a map from the topology \<open>T1\<close> to the topology \<open>T2\<close>. Then
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    52
the current \<open>continuous_on\<close> (with type classes) can be redefined as
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    64
I realized afterwards that this theory has a lot in common with \<^file>\<open>~~/src/HOL/Library/Finite_Map.thy\<close>.
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    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
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
    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
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    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
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    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
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
    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
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    87
text \<open>The total set of the product topology is the product of the total sets
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    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
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   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
07bec530f02e cleaned proofs
nipkow
parents: 71200
diff changeset
   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
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   197
lemma product_topology_trivial_iff:
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   198
   "product_topology X I = trivial_topology \<longleftrightarrow> (\<exists>i \<in> I. X i = trivial_topology)"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   199
  by (auto simp: PiE_eq_empty_iff simp flip: null_topspace_iff_trivial)
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   200
78200
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
   201
lemma topspace_product_topology_alt:
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
   202
  "topspace (product_topology X I) = {x \<in> extensional I. \<forall>i \<in> I. x i \<in> topspace(X i)}"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
   203
  by (fastforce simp: PiE_iff)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
   204
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   205
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
   206
  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
   207
  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
   208
  unfolding product_topology_def
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   209
  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
   210
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   211
proposition product_topology_open_contains_basis:
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   212
  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
   213
  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
   214
proof -
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   215
  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
   216
  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
   217
    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
   218
  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
   219
  proof induction
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   220
    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
   221
    then obtain XU XV where H:
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   222
      "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
   223
      "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
   224
      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
   225
    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
   226
    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
   227
      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
   228
    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
   229
    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
   230
      unfolding X_def using H by auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   231
    moreover have "finite (IT X)"
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   232
      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
   233
      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
   234
    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
   235
      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
   236
    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
   237
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   238
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   239
    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
   240
    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
   241
    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
   242
      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
   243
  qed auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   244
  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
   245
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   246
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   247
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
   248
   "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
   249
  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
   250
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   251
lemma openin_product_topology:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   252
   "openin (product_topology X I) =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   253
    arbitrary union_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   254
          ((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
   255
           relative_to topspace (product_topology X I))"
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   256
  by (simp add: istopology_subbase product_topology)
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   257
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   258
lemma subtopology_product_topology:
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   259
  "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
   260
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   261
  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
   262
  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
   263
  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
   264
        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
   265
    by (rule finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   266
  also have "\<dots> = finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   267
                      ((\<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
   268
                       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
                   relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   270
    apply (rule arg_cong2 [where f = "(relative_to)"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   271
     apply (rule arg_cong [where f = "(intersection_of)finite"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   272
     apply (rule ext)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   273
     apply (auto simp: relative_to_def intersection_of_def)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   274
    done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   275
  finally
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   276
  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
   277
        finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   278
          (\<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
   279
          relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   280
    by (metis finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   281
  then show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   282
  unfolding topology_eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   283
  apply clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   284
  apply (simp add: openin_product_topology flip: openin_relative_to)
71172
nipkow
parents: 70817
diff changeset
   285
  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
   286
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   287
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   288
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   289
lemma product_topology_base_alt:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   290
   "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
   291
    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
   292
    (\<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
   293
   (is "?lhs = ?rhs")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   294
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   295
  have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   296
    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
   297
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   298
    fix \<F>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   299
    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
   300
    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
   301
          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
   302
    proof (induction)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   303
      case (insert F \<F>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   304
      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
   305
        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
   306
        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
   307
        by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   308
      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
   309
        using insert by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   310
      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
   311
      show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   312
      proof (intro exI conjI)
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69722
diff changeset
   313
        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
   314
        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
   315
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   316
        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
   317
          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
   318
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   319
        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
   320
          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
   321
      qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   322
    qed (auto intro: dest: not_finite_existsD)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   323
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   324
  moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   325
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   326
    fix U :: "'a \<Rightarrow> 'b set"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   327
    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
   328
    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
   329
    show "?lhs (Pi\<^sub>E I U)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   330
      unfolding relative_to_def topspace_product_topology
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   331
    proof (intro exI conjI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   332
      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
   333
        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
   334
      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
   335
        using ope openin_subset by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   336
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   337
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   338
  ultimately show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   339
    by meson
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   340
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   341
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
   342
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
   343
  "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
   344
    (\<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
   345
                 (\<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
   346
  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
   347
  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
   348
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   349
lemma closure_of_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   350
  "(product_topology X I) closure_of (PiE I S) = PiE I (\<lambda>i. (X i) closure_of (S i))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   351
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   352
  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))
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   353
       \<longleftrightarrow> (\<forall>i \<in> I. \<forall>T. f i \<in> T \<and> openin (X i) T \<longrightarrow> S i \<inter> T \<noteq> {})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   354
    (is "?lhs = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   355
    if top: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i)" and ext:  "f \<in> extensional I" for f
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   356
  proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   357
    assume L[rule_format]: ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   358
    show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   359
    proof clarify
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   360
      fix i T
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   361
      assume "i \<in> I" "f i \<in> T" "openin (X i) T" "S i \<inter> T = {}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   362
      then have "openin (product_topology X I) ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> {x. x i \<in> T})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   363
        by (force simp: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   364
      then show "False"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   365
        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>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   366
        by (auto simp: top ext PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   367
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   368
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   369
    assume R [rule_format]: ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   370
    show ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   371
    proof (clarsimp simp: openin_product_topology union_of_def arbitrary_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   372
      fix \<U> U
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   373
      assume
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   374
        \<U>: "\<U> \<subseteq> Collect
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   375
           (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
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   376
            (\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   377
        "f \<in> U" "U \<in> \<U>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   378
      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)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   379
                 relative_to (\<Pi>\<^sub>E i\<in>I. topspace (X i))) U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   380
        by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   381
      with \<open>f \<in> U\<close> \<open>U \<in> \<U>\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   382
      obtain \<T> where "finite \<T>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   383
        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}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   384
        and "topspace (product_topology X I) \<inter> \<Inter> \<T> \<subseteq> U" "f \<in> topspace (product_topology X I) \<inter> \<Inter> \<T>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   385
        apply (clarsimp simp add: relative_to_def intersection_of_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   386
        apply (rule that, auto dest!: subsetD)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   387
        done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   388
      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"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   389
        by (auto simp: PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   390
      have *: "f i \<in> topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   391
             \<and> openin (X i) (topspace (X i) \<inter> \<Inter>{U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   392
        if "i \<in> I" for i
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   393
      proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   394
        have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   395
        proof (rule finite_vimageI [OF \<open>finite \<T>\<close>])
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   396
          show "inj (\<lambda>U. {x. x i \<in> U})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   397
            by (auto simp: inj_on_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   398
        qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   399
        then have fin: "finite {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   400
          by (rule rev_finite_subset) auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   401
        have "openin (X i) (\<Inter> (insert (topspace (X i)) {U. openin (X i) U \<and> {x. x i \<in> U} \<in> \<T>}))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   402
          by (rule openin_Inter) (auto simp: fin)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   403
        then show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   404
          using \<open>f \<in> \<Inter> \<T>\<close> by (fastforce simp: that top)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   405
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   406
      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>}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   407
      have "\<forall>i \<in> I. \<exists>x. x \<in> S i \<inter> \<Phi> i"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   408
        using R [OF _ *] unfolding \<Phi>_def by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   409
      then obtain \<theta> where \<theta> [rule_format]: "\<forall>i \<in> I. \<theta> i \<in> S i \<inter> \<Phi> i"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   410
        by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   411
      show "\<exists>y\<in>Pi\<^sub>E I S. \<exists>x\<in>\<U>. y \<in> x"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   412
      proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   413
        show "\<exists>U \<in> \<U>. (\<lambda>i \<in> I. \<theta> i) \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   414
        proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   415
          have "restrict \<theta> I \<in> PiE I (topspace \<circ> X) \<inter> \<Inter>\<T>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   416
            using \<T> by (fastforce simp: \<Phi>_def PiE_def dest: \<theta>)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   417
          then show "restrict \<theta> I \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   418
            using subU by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   419
        qed (rule \<open>U \<in> \<U>\<close>)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   420
      next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   421
        show "(\<lambda>i \<in> I. \<theta> i) \<in> Pi\<^sub>E I S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   422
          using \<theta> by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   423
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   424
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   425
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   426
  show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   427
    apply (simp add: * closure_of_def PiE_iff set_eq_iff cong: conj_cong)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   428
    by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   429
qed
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   430
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   431
corollary closedin_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   432
   "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
   433
  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
   434
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   435
corollary closedin_product_topology_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   436
   "f \<in> extensional I \<Longrightarrow> closedin (product_topology X I) {f} \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) {f i})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   437
  using PiE_singleton closedin_product_topology [of X I]
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   438
  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
   439
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   440
lemma product_topology_empty:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   441
   "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
   442
  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
   443
  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
   444
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   445
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
   446
  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
   447
  by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   448
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   449
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   450
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
   451
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
   452
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
   453
  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
   454
  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
   455
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   456
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   457
    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
   458
    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
   459
    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
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   460
      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
   461
      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
   462
    have **: "(\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   463
      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
   464
    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
   465
      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
   466
      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
   467
      apply (subst *)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   468
      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
   469
  }
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
   470
  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
   471
    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
   472
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   473
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
   474
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
   475
  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
   476
          "\<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
   477
  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
   478
unfolding product_topology_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   479
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
   480
  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
   481
  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
   482
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   483
  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
   484
  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
   485
  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
   486
    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
   487
  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
   488
    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
   489
  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
   490
    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
   491
  also have "... = (\<Inter>i\<in>J. (\<lambda>x. f x i)-`(X i) \<inter> topspace T1) \<inter> (topspace T1)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   492
    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
   493
  also have "openin T1 (...)"
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   494
    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
   495
  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
   496
next
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   497
  have "f \<in> topspace T1 \<rightarrow> topspace (product_topology T I)"
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   498
    using assms continuous_map_funspace by (force simp: Pi_iff)
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   499
  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)}}"
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   500
    by (fastforce simp add: product_topology_def Pi_iff)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   501
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   502
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
   503
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
   504
  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
   505
  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
   506
        "\<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
   507
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   508
  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
   509
  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
   510
  also have "continuous_map T1 (T i) (...)"
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   511
    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
   512
  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
   513
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   514
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   515
  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
   516
  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
   517
    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
   518
  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
   519
    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
   520
  then show "f x i = undefined"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   521
    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
   522
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   523
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   524
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
   525
  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
   526
  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
   527
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
   528
  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
   529
  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
   530
  then show "continuous_map (product_topology T I) (T i) (\<lambda>x. restrict x J i)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   531
    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
   532
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   533
  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
   534
  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
   535
    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
   536
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   537
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   538
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   539
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
   540
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   541
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
   542
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   543
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   544
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
   545
  "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
   546
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   547
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
   548
  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
   549
    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
   550
  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
   551
    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
   552
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
   553
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   554
end
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   555
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   556
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
   557
  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
   558
  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
   559
  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
   560
    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
   561
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   562
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
   563
  "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
   564
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
   565
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   566
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
   567
  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
   568
  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
   569
  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
   570
proof -
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   571
  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
   572
  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
   573
  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
   574
    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
   575
  then have "open (Pi\<^sub>E UNIV X)"
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   576
    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
   577
  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
   578
    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
   579
  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
   580
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   581
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   582
text \<open>The results proved in the general situation of products of possibly different
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   583
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
   584
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   585
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
   586
  "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
   587
  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
   588
    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
   589
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
   590
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
   591
  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
   592
  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
   593
  shows "continuous_on S f"
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   594
  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
   595
      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
   596
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
   597
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
   598
  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
   599
  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
   600
  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
   601
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
   602
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
   603
  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
   604
  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
   605
  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
   606
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   607
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
   608
  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
   609
  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
   610
  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
   611
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
   612
  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
   613
    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
   614
  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
   615
    by (intro continuous_intros) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   616
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   617
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   618
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
   619
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   620
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
   621
of product spaces, but they have more to do with countability and could
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   622
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
   623
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   624
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
   625
  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
   626
  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
   627
  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
   628
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   629
  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
   630
                  \<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
   631
    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
   632
  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
   633
  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
   634
    case 0
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   635
    have "{x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>(0::nat). x i = a)} = {(\<lambda>i. a)}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   636
      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
   637
    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
   638
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   639
    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
   640
    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
   641
      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
   642
    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
   643
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   644
      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
   645
      have "f (x(N:=a), x N) = x"
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   646
        unfolding f_def by auto
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   647
      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
   648
        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
   649
      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
   650
        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
   651
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   652
    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
   653
      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
   654
    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
   655
      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
   656
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   657
  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
   658
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   659
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   660
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
   661
  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
   662
  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
   663
  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
   664
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   665
  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
   666
  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
   667
  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
   668
  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
   669
  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
   670
        \<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
   671
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   672
    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
   673
    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
   674
    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
   675
    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
   676
    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
   677
      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
   678
    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
   679
      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
   680
    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
   681
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   682
    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
   683
      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
   684
    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
   685
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   686
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   687
  then show ?thesis
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   688
    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
   689
    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
   690
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   691
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   692
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
   693
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   694
  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
   695
  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
   696
    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
   697
  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
   698
                                                  "\<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
   699
                                                  "\<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
   700
    by metis
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   701
  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
   702
  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
   703
  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
   704
  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
   705
    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
   706
  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
   707
  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
   708
    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
   709
  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
   710
  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
   711
    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
   712
  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
   713
    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
   714
  ultimately have "countable K" by auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   715
  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
   716
    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
   717
  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
   718
    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
   719
  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
   720
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   721
    have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   722
      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
   723
    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
   724
    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
   725
      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
   726
    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
   727
                           "\<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
   728
                           "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
   729
                           "(\<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
   730
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   731
    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
   732
    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
   733
    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
   734
      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
   735
    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
   736
    proof (cases "i \<in> I")
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   737
      case True
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   738
      then show ?thesis  
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   739
        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
   740
    next
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   741
      case False
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   742
      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
   743
    qed
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   744
    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
   745
      unfolding Y_def by auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   746
    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
   747
      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
   748
    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
   749
      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
   750
    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
   751
      using "**" by auto
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   752
    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
   753
      by (metis H(4) PiE_mono subset_trans)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   754
    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
   755
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   756
  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
   757
    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
   758
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   759
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   760
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
   761
  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
   762
          topological_basis K \<and> countable K \<and>
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
   763
          (\<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
   764
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   765
  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
   766
    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
   767
  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
   768
  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
   769
  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
   770
    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
   771
  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
   772
    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
   773
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   774
  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
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
   775
  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
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   776
    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
   777
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   778
  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
   779
    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
   780
  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
   781
    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
   782
  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
   783
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   784
  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
   785
  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
   786
    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
   787
    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
   788
      unfolding open_fun_def by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   789
    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
   790
    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
   791
                           "\<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
   792
                           "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
   793
                           "(\<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
   794
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   795
    then have "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
   796
    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
   797
    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
   798
    have *: "\<exists>y. y \<in> B2 \<and> y \<subseteq> X i \<and> x i \<in> y" for i
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   799
      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
   800
    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
   801
      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
   802
    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
   803
      unfolding Y_def by auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   804
    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
   805
      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
   806
    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
   807
      unfolding K_def by auto
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   808
    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
   809
      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
   810
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   811
    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
   812
    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
   813
      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
   814
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   815
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   816
  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
   817
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   818
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   819
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
   820
proof
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   821
  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
   822
    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
   823
    by auto
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   824
qed
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   825
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   826
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
   827
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   828
theorem Alexander_subbase:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   829
  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
   830
    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
   831
  shows "compact_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   832
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   833
  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
   834
    by (simp flip: X)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   835
  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
   836
    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
   837
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   838
    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
   839
    have 1: "\<A> \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   840
      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
   841
    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
   842
      unfolding \<A>_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   843
    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
   844
      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
   845
        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
   846
      have "\<C> \<subseteq> \<A>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   847
        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
   848
      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
   849
        by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   850
      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
   851
      proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   852
        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
   853
          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
   854
        then show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   855
          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
   856
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   857
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   858
    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
   859
      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
   860
    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
   861
                        \<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
   862
           \<Longrightarrow> \<W> = \<K>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   863
      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
   864
      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
   865
      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
   866
    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
   867
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   868
      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
   869
        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
   870
      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
   871
        by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   872
      then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   873
        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
   874
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   875
    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
   876
      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
   877
    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
   878
      by (metis openin_subset)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   879
    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
   880
      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
   881
    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
   882
      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
   883
    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
   884
      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
   885
      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
   886
    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
   887
      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
   888
    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
   889
      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
   890
    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
   891
      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
   892
    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
   893
      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
   894
    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
   895
      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
   896
    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
   897
    proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   898
      fix b
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   899
      assume "b \<in> \<B>'"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   900
      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
   901
      proof (rule *)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   902
        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
   903
          using that
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   904
        proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   905
          have "b \<in> \<B>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   906
            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
   907
          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
   908
            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
   909
          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
   910
            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
   911
          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
   912
            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
   913
            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
   914
            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
   915
            done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   916
          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
   917
            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
   918
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   919
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   920
        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
   921
          using top by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   922
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   923
        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
   924
        proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   925
          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
   926
            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
   927
          then show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   928
            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
   929
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   930
      qed auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   931
      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
   932
        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
   933
        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
   934
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   935
    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
   936
      by metis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   937
    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
   938
    show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   939
    proof (rule non)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   940
      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
   941
        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
   942
      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
   943
        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
   944
      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
   945
      show "?\<D> \<subseteq> \<K>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   946
        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
   947
      show "finite ?\<D>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   948
        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
   949
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   950
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   951
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   952
    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
   953
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   954
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   955
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   956
corollary Alexander_subbase_alt:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   957
  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
   958
  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
   959
   and X: "topology
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   960
              (arbitrary union_of
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   961
                   (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
   962
 shows "compact_space X"
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 "topspace X = U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   965
    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
   966
  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
   967
    unfolding relative_to_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   968
    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
   969
  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
   970
    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
   971
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   972
    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
   973
      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
   974
    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
   975
      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
   976
    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
   977
      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
   978
    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
   979
      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
   980
    then show ?thesis
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   981
      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
   982
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   983
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   984
    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
   985
     apply (simp flip: X)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   986
     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
   987
    apply (blast intro: *)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   988
    done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   989
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   990
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   991
proposition continuous_map_componentwise:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   992
   "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
   993
    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
   994
    (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
   995
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
   996
  case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   997
  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
   998
    by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   999
  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
  1000
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1001
    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
  1002
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1003
      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
  1004
        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
  1005
        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
  1006
        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
  1007
        done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1008
      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
  1009
        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
  1010
      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
  1011
        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
  1012
      ultimately show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1013
        by metis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1014
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1015
    with that
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1016
    show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1017
      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
  1018
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1019
  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
  1020
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1021
    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
  1022
      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
  1023
    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
  1024
      by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1025
    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
  1026
      by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1027
    show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1028
      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
  1029
    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
  1030
      fix \<T>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1031
      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
  1032
                  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
  1033
      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
  1034
      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
  1035
        fix S T
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1036
        assume "T \<in> \<T>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1037
        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
  1038
          "\<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
  1039
          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
  1040
        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
  1041
          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
  1042
          unfolding 3
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1043
          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
  1044
          done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1045
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1046
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1047
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1048
  ultimately show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1049
    by metis
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1050
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
  1051
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1052
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1053
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
  1054
   "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
  1055
  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
  1056
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1057
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
  1058
   "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
  1059
  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
  1060
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1061
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
  1062
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1063
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
  1064
  assumes "i \<in> I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1065
  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
  1066
  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
  1067
proof clarify
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1068
  fix \<V>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1069
  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
  1070
               (finite intersection_of
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1071
                (\<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
  1072
                           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
  1073
  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
  1074
  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
  1075
    fix S V
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1076
    assume "V \<in> \<V>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1077
    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
  1078
      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
  1079
      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
  1080
      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
  1081
      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
  1082
    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
  1083
    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
  1084
      fix x f
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1085
      assume "f \<in> V"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1086
      let ?T = "{a \<in> topspace(Y i).
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1087
                   (\<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
  1088
      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
  1089
      proof (intro exI conjI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1090
        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
  1091
        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
  1092
          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
  1093
          proof (cases "k=i")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1094
            case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1095
            then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1096
              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
  1097
          next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1098
            case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1099
            then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1100
              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
  1101
          qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1102
          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
  1103
                  (\<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
  1104
            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
  1105
        next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1106
          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
  1107
            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
  1108
          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
  1109
                         if "\<F> \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1110
          proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1111
            show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1112
            proof (rule openin_Inter)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1113
              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
  1114
                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
  1115
                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
  1116
                using subsetD [OF \<F>]
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1117
                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
  1118
            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
  1119
          qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1120
          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
  1121
            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
  1122
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1123
      next
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1124
        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
  1125
          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
  1126
        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
  1127
          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
  1128
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1129
        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
  1130
          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
  1131
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1132
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1133
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1134
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1135
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1136
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
  1137
  assumes "i \<in> I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1138
  shows "(retraction_map (product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1139
         ((product_topology X I) = trivial_topology) \<longrightarrow> (X i) = trivial_topology)"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1140
  (is "?lhs = ?rhs")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1141
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1142
  assume ?lhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1143
  then show ?rhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1144
    using retraction_imp_surjective_map
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1145
    by (metis image_empty subtopology_eq_discrete_topology_empty)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1146
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1147
  assume R: ?rhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1148
  show ?lhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1149
  proof (cases "(product_topology X I) = trivial_topology")
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1150
    case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1151
    then show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1152
      using R by (auto simp: retraction_map_def retraction_maps_def)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1153
  next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1154
    case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1155
    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
  1156
      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
  1157
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1158
      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
  1159
        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
  1160
      show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1161
        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
  1162
        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
  1163
    qed
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1164
    with  \<open>i \<in> I\<close> False assms show ?thesis
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1165
      by (auto simp: retraction_map_def retraction_maps_def simp flip: null_topspace_iff_trivial)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1166
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1167
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1168
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
  1169
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
  1170
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
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
  1172
  "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
  1173
        PiE I S = {} \<or>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1174
        finite {i \<in> I. S i \<noteq> topspace (X i)} \<and> (\<forall>i \<in> I. openin (X i) (S 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
  1175
  (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
  1176
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
  1177
  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
  1178
  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
  1179
  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
  1180
    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
  1181
    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
  1182
    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
  1183
      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
  1184
    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
  1185
      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
  1186
      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
  1187
      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
  1188
    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
  1189
      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
  1190
    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
  1191
    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
  1192
      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
  1193
        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
  1194
        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
  1195
        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
  1196
    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
  1197
      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
  1198
      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
  1199
      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
  1200
        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
  1201
        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
  1202
        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
  1203
        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
  1204
        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
  1205
    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
  1206
  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
  1207
    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
  1208
    then show ?lhs
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1209
      unfolding openin_product_topology
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1210
      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
  1211
  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
  1212
  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
  1213
    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
  1214
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
  1215
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
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
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
  1218
   "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
  1219
  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
  1220
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
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
  1222
   "compact_space(product_topology X I) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1223
    (product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. compact_space(X 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
  1224
    (is "?lhs = ?rhs")
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1225
proof (cases "(product_topology X I) = trivial_topology")
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
  1226
  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
  1227
  then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1228
    by (auto simp flip: null_topspace_iff_trivial)
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
  1229
  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
  1230
  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
  1231
    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
  1232
    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
  1233
    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
  1234
      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
  1235
      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
  1236
      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
  1237
        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
  1238
      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
  1239
      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
  1240
        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
  1241
      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
  1242
        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
  1243
      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
  1244
        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
  1245
    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
  1246
  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
  1247
    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
  1248
    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
  1249
    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
  1250
      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
  1251
      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
  1252
        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
  1253
    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
  1254
      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
  1255
      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
  1256
        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
  1257
      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
  1258
        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
  1259
      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
  1260
        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
  1261
        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
  1262
        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
  1263
        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
  1264
          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
  1265
            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
  1266
        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
  1267
          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
  1268
          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
  1269
          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
  1270
          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
  1271
          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
  1272
            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
  1273
            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
  1274
                   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
  1275
              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
  1276
            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
  1277
                            \<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
  1278
              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
  1279
            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
  1280
              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
  1281
            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
  1282
              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
  1283
            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
  1284
              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
  1285
              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
  1286
          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
  1287
            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
  1288
            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
  1289
              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
  1290
            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
  1291
              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
  1292
            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
  1293
              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
  1294
            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
  1295
              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
  1296
            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
  1297
              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
  1298
          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
  1299
        qed (simp add: product_topology)
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1300
      qed simp
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
  1301
    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
  1302
  qed
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1303
qed 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
  1304
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
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
  1306
   "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
  1307
        PiE I S = {} \<or> (\<forall>i \<in> I. compactin (X i) (S i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1308
  by (fastforce simp add: compactin_subspace subtopology_product_topology compact_space_product_topology
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1309
      subset_PiE product_topology_trivial_iff subtopology_trivial_iff)
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
  1310
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
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
  1312
   "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
  1313
        \<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
  1314
  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
  1315
  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
  1316
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
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
  1318
   "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
  1319
  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
  1320
  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
  1321
  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
  1322
     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
  1323
  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
  1324
   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
  1325
  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
  1326
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1327
subsection\<open>Relationship with connected spaces, paths, etc.\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1328
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1329
proposition connected_space_product_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1330
   "connected_space(product_topology X I) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1331
    (\<exists>i \<in> I. X i = trivial_topology) \<or> (\<forall>i \<in> I. connected_space(X i))"
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1332
  (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1333
proof (cases ?eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1334
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1335
  moreover have "?lhs = ?rhs"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1336
  proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1337
    assume ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1338
    moreover
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1339
    have "connectedin(X i) (topspace(X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1340
      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
  1341
    proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1342
      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
  1343
        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
  1344
      show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1345
        using connectedin_continuous_map_image [OF cm ci] \<open>i \<in> I\<close>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1346
        by (simp add: False image_projection_PiE PiE_eq_empty_iff)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1347
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1348
    ultimately show ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1349
      by (meson connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1350
  next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1351
    assume cs [rule_format]: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1352
    have False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1353
      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
  1354
        and U: "openin (product_topology X I) U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1355
        and V: "openin (product_topology X I) V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1356
        and "U \<noteq> {}" "V \<noteq> {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1357
      for U V
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 f where "f \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1360
        using \<open>U \<noteq> {}\<close> by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1361
      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
  1362
        using U openin_subset by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1363
      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
  1364
        using U V openin_subset by blast+
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1365
      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
  1366
      proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1367
        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
  1368
            (\<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
  1369
          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
  1370
        then obtain \<T> where "finite \<T>"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1371
          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
  1372
          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
  1373
          and ftop: "f \<in> topspace (product_topology X I)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1374
          and fint: "f \<in> \<Inter> \<T>"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1375
          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
  1376
        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
  1377
        obtain L where "finite L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1378
          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
  1379
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1380
          show "finite ?L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1381
          proof (rule finite_Union)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1382
            fix M
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1383
            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
  1384
            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
  1385
              by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1386
            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
  1387
              using t by meson
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1388
            then have "f j \<in> V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1389
              using \<open>C \<in> \<T>\<close> fint by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1390
            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
  1391
              using that
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1392
              apply (clarsimp simp add: set_eq_iff)
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1393
              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
  1394
              done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1395
            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
  1396
              using Ceq by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1397
            then show "finite M"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1398
              using C finite_subset by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1399
          qed (use \<open>finite \<T>\<close> in blast)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1400
        next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1401
          fix i U
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1402
          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
  1403
          then show "i \<in> ?L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1404
            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
  1405
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1406
        show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1407
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1408
          fix h
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1409
          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
  1410
          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
  1411
          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
  1412
            unfolding g_def using f h by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1413
          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
  1414
            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
  1415
          ultimately have "g \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1416
            using subU by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1417
          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
  1418
            using that
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1419
          proof (induction arbitrary: h)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1420
            case empty
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1421
            then show ?case
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1422
              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
  1423
          next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1424
            case (insert i M)
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1425
            define f where "f \<equiv> h(i:=g i)"
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1426
            have fin: "f \<in> PiE I (topspace \<circ> X)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1427
              unfolding f_def using gin insert.prems(1) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1428
            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
  1429
              unfolding f_def using insert.prems(2) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1430
            have "f \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1431
              using insert.IH [OF fin subM] .
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1432
            show ?case
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1433
            proof (cases "h \<in> V")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1434
              case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1435
              show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1436
              proof (cases "i \<in> I")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1437
                case True
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1438
                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
  1439
                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
  1440
                have False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1441
                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
  1442
                  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
  1443
                    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
  1444
                  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
  1445
                    using \<open>i \<in> I\<close> insert.prems(1)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1446
                    by (auto simp: continuous_map_componentwise extensional_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1447
                  show "openin (X i) ?U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1448
                    by (rule openin_continuous_map_preimage [OF cm U])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1449
                  show "openin (X i) ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1450
                    by (rule openin_continuous_map_preimage [OF cm V])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1451
                  show "topspace (X i) \<subseteq> ?U \<union> ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1452
                  proof clarsimp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1453
                    fix x
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1454
                    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
  1455
                    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
  1456
                    show "h(i:=x) \<in> U"
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1457
                      by (force dest: subsetD [where c="h(i:=x)"])
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1458
                  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1459
                  show "?U \<inter> ?V = {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1460
                    using disj by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1461
                  show "?U \<noteq> {}"
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1462
                    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
  1463
                  show "?V \<noteq> {}"
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1464
                    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
  1465
                qed
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
              next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1468
                case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1469
                show ?thesis
76821
337c2265d8a2 Further cleaning up of messy proofs
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1470
                  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
  1471
              qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1472
            next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1473
              case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1474
              then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1475
                using subUV insert.prems(1) by auto
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
          qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1478
          then show "h \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1479
            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
  1480
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1481
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1482
      ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1483
        using disj inf_absorb2 \<open>V \<noteq> {}\<close> by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1484
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1485
    then show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1486
      unfolding connected_space_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1487
      by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1488
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1489
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1490
    by simp
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1491
qed (metis connected_space_trivial_topology product_topology_trivial_iff)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1492
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1493
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1494
lemma connectedin_PiE:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1495
   "connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1496
        PiE I S = {} \<or> (\<forall>i \<in> I. connectedin (X i) (S i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1497
  by (auto simp: connectedin_def subtopology_product_topology connected_space_product_topology subset_PiE 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1498
      PiE_eq_empty_iff subtopology_trivial_iff)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1499
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1500
lemma path_connected_space_product_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1501
   "path_connected_space(product_topology X I) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1502
    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
  1503
 (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1504
proof (cases ?eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1505
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1506
  moreover have "?lhs = ?rhs"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1507
  proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1508
    assume L: ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1509
    show ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1510
    proof (clarsimp simp flip: path_connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1511
      fix i :: "'a"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1512
      assume "i \<in> I"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1513
      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
  1514
        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
  1515
      show "path_connectedin (X i) (topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1516
        using path_connectedin_continuous_map_image [OF cm L [unfolded path_connectedin_topspace [symmetric]]]
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1517
        by (metis \<open>i \<in> I\<close> False retraction_imp_surjective_map retraction_map_product_projection topspace_discrete_topology)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1518
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1519
  next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1520
    assume R [rule_format]: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1521
    show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1522
      unfolding path_connected_space_def topspace_product_topology
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1523
    proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1524
      fix x y
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1525
      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
  1526
      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
  1527
        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
  1528
      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
  1529
        by metis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1530
      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
  1531
        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
  1532
        apply (force simp: pathin_def continuous_map_componentwise)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1533
        done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1534
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1535
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1536
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1537
    by simp
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1538
next
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1539
qed (force simp: path_connected_space_topspace_empty iff: null_topspace_iff_trivial)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1540
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1541
lemma path_connectedin_PiE:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1542
   "path_connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1543
    PiE I S = {} \<or> (\<forall>i \<in> I. path_connectedin (X i) (S i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1544
  by (fastforce simp add: path_connectedin_def subtopology_product_topology path_connected_space_product_topology subset_PiE PiE_eq_empty_iff topspace_subtopology_subset)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1545
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1546
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
  1547
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1548
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
  1549
  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
  1550
  shows "quotient_map(product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1551
           ((product_topology X I) = trivial_topology \<longrightarrow> (X i) = trivial_topology)"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1552
  by (metis (no_types) assms image_is_empty null_topspace_iff_trivial quotient_imp_surjective_map 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1553
      retraction_imp_quotient_map 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
  1554
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1555
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
  1556
  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
  1557
  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
  1558
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1559
  have "quotient_map (product_topology X I) (X i) (\<lambda>x. x i)"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1560
    using assms  by (metis (full_types) discrete_topology_unique empty_not_insert 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1561
             product_topology_trivial_iff quotient_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
  1562
  moreover
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1563
  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
  1564
    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
  1565
  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
  1566
    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
  1567
    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
  1568
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1569
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1570
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
  1571
  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
  1572
    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
  1573
                      \<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
  1574
               (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
  1575
    and PQ:  "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1576
  shows "(\<exists>i\<in>I. (X i) = trivial_topology) \<or> (\<forall>i \<in> I. Q(X i))"
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1577
proof -
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1578
  have "Q(X i)" if "\<forall>i\<in>I. (X i) \<noteq> trivial_topology" "i \<in> I" for i
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1579
  proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1580
    from that obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" 
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1581
      by (meson null_topspace_iff_trivial PiE_eq_empty_iff ex_in_conv)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1582
    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
  1583
      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})"]
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1584
      by (force simp add: subtopology_product_topology)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1585
    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
  1586
      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
  1587
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  1588
  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
  1589
qed
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1590
78200
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1591
subsection \<open>Limits\<close>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1592
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1593
text \<open>The original HOL Light proof was a mess, yuk\<close>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1594
lemma limitin_componentwise:
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1595
   "limitin (product_topology X I) f l F \<longleftrightarrow>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1596
        l \<in> extensional I \<and>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1597
        eventually (\<lambda>a. f a \<in> topspace(product_topology X I)) F \<and>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1598
        (\<forall>i \<in> I. limitin (X i) (\<lambda>c. f c i) (l i) F)"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1599
    (is "?L \<longleftrightarrow> _ \<and> ?R1 \<and> ?R2")
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1600
proof (cases "l \<in> extensional I")
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1601
  case l: True
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1602
  show ?thesis
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1603
  proof (cases "\<forall>i\<in>I. l i \<in> topspace (X i)")
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1604
    case True
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1605
    have ?R1 if ?L
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1606
      by (metis limitin_subtopology subtopology_topspace that)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1607
    moreover
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1608
    have ?R2 if ?L
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1609
      unfolding limitin_def
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1610
    proof (intro conjI strip)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1611
      fix i U
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1612
      assume "i \<in> I" and U: "openin (X i) U \<and> l i \<in> U"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1613
      then have "openin (product_topology X I) ({y. y i \<in> U} \<inter> topspace (product_topology X I))"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1614
        unfolding openin_product_topology arbitrary_union_of_relative_to [symmetric]
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1615
        apply (simp add: relative_to_def topspace_product_topology_alt)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1616
        by (smt (verit, del_insts) Collect_cong arbitrary_union_of_inc finite_intersection_of_inc inf_commute)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1617
      moreover have "l \<in> {y. y i \<in> U} \<inter> topspace (product_topology X I)"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1618
        using U True l by (auto simp: extensional_def)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1619
      ultimately have "eventually (\<lambda>x. f x \<in> {y. y i \<in> U} \<inter> topspace (product_topology X I)) F"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1620
        by (metis limitin_def that)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1621
      then show "\<forall>\<^sub>F x in F. f x i \<in> U"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1622
        by (simp add: eventually_conj_iff)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1623
    qed (use True in auto)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1624
    moreover
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1625
    have ?L if R1: ?R1 and R2: ?R2
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1626
      unfolding limitin_def openin_product_topology all_union_of imp_conjL arbitrary_def
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1627
    proof (intro conjI strip)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1628
      show l: "l \<in> topspace (product_topology X I)"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1629
        by (simp add: PiE_iff True l)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1630
      fix \<V>
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1631
      assume "\<V> \<subseteq> Collect (finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (X i) U)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1632
                   relative_to topspace (product_topology X I))"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1633
        and "l \<in> \<Union> \<V>"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1634
      then obtain \<W> where "finite \<W>" and \<W>X: "\<forall>X\<in>\<W>. l \<in> X"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1635
               and \<W>: "\<And>C. C \<in> \<W> \<Longrightarrow> C \<in> {{x. x i \<in> U} |i U. i \<in> I \<and> openin (X i) U}" 
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1636
               and \<W>\<V>: "topspace (product_topology X I) \<inter> \<Inter> \<W> \<in> \<V>"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1637
        by (fastforce simp: intersection_of_def relative_to_def subset_eq)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1638
      have "\<forall>\<^sub>F x in F. f x \<in> topspace (product_topology X I) \<inter> \<Inter>\<W>"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1639
      proof -
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1640
        have "\<And>W. W \<in> {{x. x i \<in> U} | i U. i \<in> I \<and> openin (X i) U} \<Longrightarrow> W \<in> \<W> \<Longrightarrow> \<forall>\<^sub>F x in F. f x \<in> W"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1641
          using \<W>X R2 by (auto simp: limitin_def)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1642
        with \<W> have "\<forall>\<^sub>F x in F. \<forall>W\<in>\<W>. f x \<in> W"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1643
          by (simp add: \<open>finite \<W>\<close> eventually_ball_finite)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1644
        with R1 show ?thesis
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1645
          by (simp add: eventually_conj_iff)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1646
      qed
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1647
      then show "\<forall>\<^sub>F x in F. f x \<in> \<Union>\<V>"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1648
        by (smt (verit, ccfv_threshold) \<W>\<V> UnionI eventually_mono)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1649
    qed
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1650
    ultimately show ?thesis
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1651
      using l by blast
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1652
  next
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1653
    case False
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1654
    then show ?thesis
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1655
      by (metis PiE_iff limitin_def topspace_product_topology)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1656
   qed
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1657
next
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1658
  case False
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1659
  then show ?thesis
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1660
    by (simp add: limitin_def PiE_iff)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1661
qed
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 76838
diff changeset
  1662
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1663
end