src/HOL/Analysis/Function_Topology.thy
author paulson <lp15@cam.ac.uk>
Wed, 17 Apr 2019 17:48:28 +0100
changeset 70178 4900351361b0
parent 70136 f03a01a18c6e
child 70817 dd675800469d
permissions -rw-r--r--
Lindelöf spaces and supporting material
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
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
     6
imports Topology_Euclidean_Space Abstract_Limits
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
     7
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
     8
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
     9
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
    10
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
    11
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
    12
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
    13
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    14
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
    15
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
    16
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
    17
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
    18
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    19
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
    20
'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
    21
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
    22
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    23
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
    24
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
    25
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
    26
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
    27
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
    28
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
    29
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
    30
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
    31
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
    32
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    33
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
    34
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
    35
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
    36
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
    37
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    38
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
    39
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
    40
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
    41
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
    42
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    43
Here is an example of a reformulation using topologies. Let
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    44
@{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
    45
\<open>continuous_map T1 T2 f =
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    46
    ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    47
            \<and> (f`(topspace T1) \<subseteq> (topspace T2)))\<close>}
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    48
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
    49
the current \<open>continuous_on\<close> (with type classes) can be redefined as
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    50
@{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
    51
    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
    52
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
    53
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
    54
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
    55
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
    56
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    57
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
    58
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
    59
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
    60
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    61
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
    62
\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    63
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
    64
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
    65
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    66
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
    67
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
    68
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
    69
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
    70
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
    71
(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
    72
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
    73
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    74
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
    75
\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    76
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
    77
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
    78
  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
    79
    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
    80
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
    81
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
    82
  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
    83
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    84
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
    85
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
    86
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    87
proposition product_topology:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    88
  "product_topology X I =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    89
   topology
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    90
    (arbitrary union_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    91
       ((finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    92
        (\<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
    93
        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
    94
  (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
    95
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    96
  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
    97
  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
    98
  proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    99
    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
   100
      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
   101
    proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   102
      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
   103
        using \<U> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   104
      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
   105
        by metis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   106
      obtain U where "U \<in> \<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   107
        using \<open>\<U> \<noteq> {}\<close> by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   108
      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
   109
      show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   110
      proof (intro conjI exI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   111
        show "finite (\<Union>U\<in>\<U>. ?F U)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   112
          using Y \<open>finite' \<U>\<close> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   113
        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
   114
        proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   115
          have *: "f \<in> U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   116
            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
   117
              and "\<forall>i\<in>I. f i \<in> topspace (X i)" and "f \<in> extensional I" for f U
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   118
          proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   119
            have "Pi\<^sub>E I (Y U) = U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   120
              using Y \<open>U \<in> \<U>\<close> by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   121
            then show "f \<in> U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   122
              using that unfolding PiE_def Pi_def by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   123
          qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   124
          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
   125
            by (auto simp: PiE_iff *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   126
          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
   127
            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
   128
        qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   129
      qed (use Y openin_subset in \<open>blast+\<close>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   130
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   131
    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
   132
      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
   133
    proof (cases "\<U>={}")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   134
      case True
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   135
      then show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   136
        apply (rule_tac x="{?TOP}" in exI, simp)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   137
        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
   138
        apply force
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   139
        done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   140
    next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   141
      case False
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   142
      then obtain U where "U \<in> \<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   143
        by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   144
      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
   145
        using \<U> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   146
      then obtain J Y where
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   147
        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
   148
        by metis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   149
      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
   150
      show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   151
      proof (intro conjI exI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   152
        show "finite (?G ` \<U>)" "?G ` \<U> \<noteq> {}"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   153
          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
   154
        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
   155
          using Y by force
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   156
        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
   157
          apply clarsimp
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   158
          apply (rule_tac x=" (\<lambda>i. if i = J U then Y U else topspace (X i))" in exI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   159
          apply (auto simp: *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   160
          done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   161
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   162
        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
   163
        proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   164
          have "(\<Pi>\<^sub>E i\<in>I. if i = J U then Y U else topspace (X i)) \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   165
            apply (clarsimp simp: PiE_iff Ball_def all_conj_distrib split: if_split_asm)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   166
            using Y \<open>U \<in> \<U>\<close> openin_subset by blast+
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   167
          then have "(\<Inter>U\<in>\<U>. ?G U) \<subseteq> ?TOP"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   168
            using \<open>U \<in> \<U>\<close>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   169
            by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   170
          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
   171
            using PiE_mem Y by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   172
          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
   173
            by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   174
        qed (use Y in fastforce)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   175
      qed
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 relative_to_def intersection_of_def
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   179
      by (safe; blast dest!: 1 2)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   180
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   181
  show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   182
    unfolding product_topology_def generate_topology_on_eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   183
    apply (rule arg_cong [where f = topology])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   184
    apply (rule arg_cong [where f = "(union_of)arbitrary"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   185
    apply (force simp: *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   186
    done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   187
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   188
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   189
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
   190
  "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
   191
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   192
  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
   193
    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
   194
    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
   195
  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
   196
    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
   197
  then show "(\<Pi>\<^sub>E i\<in>I. topspace (T i)) \<subseteq> topspace (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
   198
    unfolding product_topology_def using PiE_def by (auto simp add: topology_generated_by_topspace)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   199
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   200
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   201
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
   202
  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
   203
  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
   204
  unfolding product_topology_def
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   205
  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
   206
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   207
proposition product_topology_open_contains_basis:
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   208
  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
   209
  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
   210
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   211
  have "generate_topology_on {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}} U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   212
    using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   213
  then have "\<And>x. x\<in>U \<Longrightarrow> \<exists>X. x \<in> (\<Pi>\<^sub>E i\<in>I. X i) \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> (\<Pi>\<^sub>E i\<in>I. X i) \<subseteq> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   214
  proof induction
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   215
    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
   216
    then obtain XU XV where H:
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   217
      "x \<in> Pi\<^sub>E I XU" "(\<forall>i. openin (T i) (XU i))" "finite {i. XU i \<noteq> topspace (T i)}" "Pi\<^sub>E I XU \<subseteq> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   218
      "x \<in> Pi\<^sub>E I XV" "(\<forall>i. openin (T i) (XV i))" "finite {i. XV i \<noteq> topspace (T i)}" "Pi\<^sub>E I XV \<subseteq> V"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   219
      by auto meson
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   220
    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
   221
    have "Pi\<^sub>E I X \<subseteq> Pi\<^sub>E I XU \<inter> Pi\<^sub>E I XV"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   222
      unfolding X_def by (auto simp add: PiE_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   223
    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
   224
    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
   225
      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
   226
    moreover have "finite {i. X i \<noteq> topspace (T i)}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   227
      apply (rule rev_finite_subset[of "{i. XU i \<noteq> topspace (T i)} \<union> {i. XV i \<noteq> topspace (T i)}"])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   228
      unfolding X_def using H by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   229
    moreover have "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
   230
      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
   231
    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
   232
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   233
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   234
    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
   235
    then obtain k where "k \<in> K" "x \<in> k" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   236
    with UN have "\<exists>X. x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   237
      by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   238
    then obtain X where "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> k"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   239
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   240
    then have "x \<in> Pi\<^sub>E I X \<and> (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)} \<and> Pi\<^sub>E I X \<subseteq> (\<Union>K)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   241
      using \<open>k \<in> K\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   242
    then show ?case
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   243
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   244
  qed auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   245
  then show ?thesis using \<open>x \<in> U\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   246
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   247
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   248
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
   249
   "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
   250
  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
   251
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   252
lemma openin_product_topology:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   253
   "openin (product_topology X I) =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   254
    arbitrary union_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   255
          ((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
   256
           relative_to topspace (product_topology X I))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   257
  apply (subst product_topology)
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   258
  apply (simp add: topology_inverse' [OF istopology_subbase])
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   259
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   260
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   261
lemma subtopology_PiE:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   262
  "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
   263
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   264
  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
   265
  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
   266
  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
   267
        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
   268
    by (rule finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   269
  also have "\<dots> = finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   270
                      ((\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> (openin (X i) relative_to S i) U)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   271
                       relative_to ?X \<inter> Pi\<^sub>E I S)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   272
                   relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   273
    apply (rule arg_cong2 [where f = "(relative_to)"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   274
     apply (rule arg_cong [where f = "(intersection_of)finite"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   275
     apply (rule ext)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   276
     apply (auto simp: relative_to_def intersection_of_def)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   277
    done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   278
  finally
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   279
  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
   280
        finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   281
          (\<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
   282
          relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   283
    by (metis finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   284
  then show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   285
  unfolding topology_eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   286
  apply clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   287
  apply (simp add: openin_product_topology flip: openin_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   288
  apply (simp add: arbitrary_union_of_relative_to topspace_product_topology topspace_subtopology flip: PiE_Int)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   289
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   290
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   291
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   292
lemma product_topology_base_alt:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   293
   "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
   294
    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
   295
    (\<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
   296
   (is "?lhs = ?rhs")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   297
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   298
  have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   299
    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
   300
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   301
    fix \<F>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   302
    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
   303
    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
   304
          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
   305
    proof (induction)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   306
      case (insert F \<F>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   307
      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
   308
        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
   309
        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
   310
        by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   311
      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
   312
        using insert by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   313
      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
   314
      show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   315
      proof (intro exI conjI)
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69722
diff changeset
   316
        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
   317
        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
   318
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   319
        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
   320
          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
   321
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   322
        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
   323
          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
   324
      qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   325
    qed (auto intro: dest: not_finite_existsD)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   326
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   327
  moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   328
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   329
    fix U :: "'a \<Rightarrow> 'b set"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   330
    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
   331
    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
   332
    show "?lhs (Pi\<^sub>E I U)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   333
      unfolding relative_to_def topspace_product_topology
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   334
    proof (intro exI conjI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   335
      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
   336
        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
   337
      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
   338
        using ope openin_subset by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   339
    qed
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
  ultimately show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   342
    by meson
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   343
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   344
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
   345
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
   346
  "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
   347
    (\<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
   348
                 (\<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
   349
  unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   350
  apply safe
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   351
  apply (drule bspec; blast)+
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   352
  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
   353
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   354
lemma closure_of_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   355
  "(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
   356
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   357
  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
   358
       \<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
   359
    (is "?lhs = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   360
    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
   361
  proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   362
    assume L[rule_format]: ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   363
    show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   364
    proof clarify
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   365
      fix i T
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   366
      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
   367
      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
   368
        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
   369
      then show "False"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   370
        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
   371
        by (auto simp: top ext PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   372
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   373
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   374
    assume R [rule_format]: ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   375
    show ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   376
    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
   377
      fix \<U> U
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   378
      assume
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   379
        \<U>: "\<U> \<subseteq> Collect
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   380
           (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
   381
            (\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   382
        "f \<in> U" "U \<in> \<U>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   383
      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
   384
                 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
   385
        by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   386
      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
   387
      obtain \<T> where "finite \<T>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   388
        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
   389
        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
   390
        apply (clarsimp simp add: relative_to_def intersection_of_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   391
        apply (rule that, auto dest!: subsetD)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   392
        done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   393
      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
   394
        by (auto simp: PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   395
      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
   396
             \<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
   397
        if "i \<in> I" for i
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   398
      proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   399
        have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   400
        proof (rule finite_vimageI [OF \<open>finite \<T>\<close>])
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   401
          show "inj (\<lambda>U. {x. x i \<in> U})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   402
            by (auto simp: inj_on_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   403
        qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   404
        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
   405
          by (rule rev_finite_subset) auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   406
        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
   407
          by (rule openin_Inter) (auto simp: fin)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   408
        then show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   409
          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
   410
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   411
      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
   412
      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
   413
        using R [OF _ *] unfolding \<Phi>_def by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   414
      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
   415
        by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   416
      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
   417
      proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   418
        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
   419
        proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   420
          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
   421
            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
   422
          then show "restrict \<theta> I \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   423
            using subU by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   424
        qed (rule \<open>U \<in> \<U>\<close>)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   425
      next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   426
        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
   427
          using \<theta> by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   428
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   429
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   430
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   431
  show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   432
    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
   433
    by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   434
qed
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   435
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   436
corollary closedin_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   437
   "closedin (product_topology X I) (PiE I S) \<longleftrightarrow> PiE I S = {} \<or> (\<forall>i \<in> I. closedin (X i) (S i))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   438
  apply (simp add: PiE_eq PiE_eq_empty_iff closure_of_product_topology flip: closure_of_eq)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   439
  apply (metis closure_of_empty)
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   440
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   441
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   442
corollary closedin_product_topology_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   443
   "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
   444
  using PiE_singleton closedin_product_topology [of X I]
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   445
  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
   446
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   447
lemma product_topology_empty:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   448
   "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
   449
  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
   450
  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
   451
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   452
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
   453
  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
   454
  by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   455
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   456
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   457
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
   458
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
   459
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
   460
  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
   461
  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
   462
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   463
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   464
    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
   465
    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
   466
    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
   467
      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
   468
      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
   469
    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
   470
      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
   471
    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
   472
      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
   473
      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
   474
      apply (subst *)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   475
      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
   476
  }
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
  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
   478
    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
   479
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   480
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
   481
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
   482
  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
   483
          "\<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
   484
  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
   485
unfolding product_topology_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   486
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
   487
  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
   488
  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
   489
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   490
  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
   491
  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
   492
  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
   493
    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
   494
  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
   495
    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
   496
  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
   497
    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
   498
  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
   499
    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
   500
  also have "openin T1 (...)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   501
    apply (rule openin_INT)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   502
    apply (simp add: \<open>finite J\<close>)
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   503
    using H(2) assms(1) \<open>J \<subseteq> I\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   504
  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
   505
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   506
  show "f `topspace T1 \<subseteq> \<Union>{Pi\<^sub>E I X |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   507
    apply (subst topology_generated_by_topspace[symmetric])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   508
    apply (subst product_topology_def[symmetric])
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   509
    apply (simp only: topspace_product_topology)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   510
    apply (auto simp add: PiE_iff)
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   511
    using assms unfolding continuous_map_def by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   512
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   513
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
   514
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
   515
  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
   516
  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
   517
        "\<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
   518
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   519
  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
   520
  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
   521
  also have "continuous_map T1 (T i) (...)"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   522
    apply (rule continuous_map_compose[of _ "product_topology T I"])
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   523
    using assms \<open>i \<in> I\<close> by auto
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   524
  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
   525
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   526
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   527
  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
   528
  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
   529
    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
   530
  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
   531
    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
   532
  then show "f x i = undefined"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   533
    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
   534
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   535
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   536
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
   537
  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
   538
  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
   539
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
   540
  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
   541
  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
   542
  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
   543
    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
   544
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   545
  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
   546
  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
   547
    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
   548
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   549
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   550
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   551
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
   552
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   553
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
   554
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   555
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   556
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
   557
  "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
   558
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   559
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
   560
  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
   561
    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
   562
  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
   563
    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
   564
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
   565
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   566
end
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   567
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   568
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
   569
  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
   570
  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
   571
  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
   572
    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
   573
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   574
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
   575
  "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
   576
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
   577
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   578
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
   579
  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
   580
  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
   581
  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
   582
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   583
  define J where "J = x`I"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   584
  define V where "V = (\<lambda>y. if y \<in> J then \<Inter>{U i|i. i\<in>I \<and> x i = y} else UNIV)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   585
  define X where "X = (\<lambda>y. if y \<in> J then V y else UNIV)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   586
  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
   587
    unfolding X_def V_def using assms by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   588
  have **: "finite {i. X i \<noteq> UNIV}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   589
    unfolding X_def V_def J_def using assms(1) by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   590
  have "open (Pi\<^sub>E UNIV X)"
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   591
    by (simp add: "*" "**" open_PiE)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   592
  moreover have "Pi\<^sub>E UNIV X = {f. \<forall>i\<in>I. f (x i) \<in> U i}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   593
    apply (auto simp add: PiE_iff) unfolding X_def V_def J_def
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   594
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   595
      fix f :: "'a \<Rightarrow> 'b" and i :: 'i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   596
      assume a1: "i \<in> I"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   597
      assume a2: "\<forall>i. f i \<in> (if i \<in> x`I then if i \<in> x`I then \<Inter>{U ia |ia. ia \<in> I \<and> x ia = i} else UNIV else UNIV)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   598
      have f3: "x i \<in> x`I"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   599
        using a1 by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   600
      have "U i \<in> {U ia |ia. ia \<in> I \<and> x ia = x i}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   601
        using a1 by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   602
      then show "f (x i) \<in> U i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   603
        using f3 a2 by (meson Inter_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   604
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   605
  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
   606
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   607
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   608
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
   609
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
   610
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   611
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
   612
  "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
   613
  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
   614
    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
   615
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
   616
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
   617
  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
   618
  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
   619
  shows "continuous_on S f"
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   620
  using continuous_map_coordinatewise_then_product [of UNIV, where T = "\<lambda>i. euclidean"]
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   621
  by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   622
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
   623
lemma continuous_on_product_then_coordinatewise_UNIV:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   624
  assumes "continuous_on UNIV f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   625
  shows "continuous_on UNIV (\<lambda>x. f x i)"
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   626
  unfolding continuous_map_iff_continuous2 [symmetric]
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   627
  by (rule continuous_map_product_then_coordinatewise [where I=UNIV]) (use assms in \<open>auto simp: euclidean_product_topology\<close>)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   628
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
   629
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
   630
  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
   631
  shows "continuous_on S (\<lambda>x. f x i)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   632
proof -
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
   633
  have "continuous_on S ((\<lambda>q. q i) \<circ> f)"
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
   634
    by (metis assms continuous_on_compose continuous_on_id
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   635
        continuous_on_product_then_coordinatewise_UNIV continuous_on_subset subset_UNIV)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   636
  then show ?thesis
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   637
    by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   638
qed
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
   639
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
   640
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
   641
  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
   642
  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
   643
  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
   644
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   645
lemma continuous_map_span_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   646
  fixes B :: "'a::real_inner set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   647
  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
   648
  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
   649
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
   650
  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
   651
    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
   652
  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
   653
    by (intro continuous_intros) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   654
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
   655
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   656
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
   657
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   658
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
   659
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
   660
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
   661
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   662
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
   663
  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
   664
  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
   665
  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
   666
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   667
  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
   668
                  \<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
   669
    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
   670
  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
   671
  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
   672
    case 0
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   673
    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
   674
      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
   675
    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
   676
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   677
    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
   678
    define f::"((nat \<Rightarrow> 'a) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   679
      where "f = (\<lambda>(x, b). (\<lambda>i. if i = N then b else x i))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   680
    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
   681
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   682
      fix x assume H: "\<forall>i::nat. x i \<in> F" "\<forall>i\<ge>Suc N. x i = a"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   683
      define y where "y = (\<lambda>i. if i = N then a else x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   684
      have "f (y, x N) = x"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   685
        unfolding f_def y_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   686
      moreover have "(y, x N) \<in> {x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   687
        unfolding y_def using H \<open>a \<in> F\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   688
      ultimately show "x \<in> f`({x. (\<forall>i. x i \<in> F) \<and> (\<forall>i\<ge>N. x i = a)} \<times> F)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   689
        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
   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
    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
   692
      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
   693
    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
   694
      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
   695
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   696
  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
   697
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   698
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   699
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
   700
  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
   701
  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
   702
  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
   703
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   704
  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
   705
  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
   706
  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
   707
  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
   708
  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
   709
        \<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
   710
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   711
    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
   712
    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
   713
    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
   714
    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
   715
    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
   716
      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
   717
    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
   718
      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
   719
    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
   720
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   721
    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
   722
      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
   723
    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
   724
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   725
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   726
  then show ?thesis
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   727
    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
   728
    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
   729
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   730
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   731
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
   732
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   733
  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
   734
  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
   735
    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
   736
  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
   737
                                                  "\<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
   738
                                                  "\<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
   739
    by metis
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   740
  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
   741
  define B where "B = (\<lambda>i. (A (x i))`UNIV \<union> {UNIV})"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   742
  have "countable (B i)" for i unfolding B_def by auto
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   743
  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
   744
    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
   745
  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
   746
  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
   747
    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
   748
  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
   749
  have "countable {X. (\<forall>i. X i \<in> B i) \<and> finite {i. X i \<noteq> UNIV}}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   750
    apply (rule countable_product_event_const) using \<open>\<And>i. countable (B i)\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   751
  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
   752
    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
   753
  ultimately have "countable K" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   754
  have "x \<in> k" if "k \<in> K" for k
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   755
    using that unfolding K_def B_def apply auto using A(1) by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   756
  have "open k" if "k \<in> K" for k
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   757
    using that unfolding K_def  by (blast intro: open_B open_PiE elim: )
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   758
  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
   759
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   760
    have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   761
      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
   762
    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
   763
    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
   764
      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
   765
    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
   766
                           "\<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
   767
                           "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
   768
                           "(\<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
   769
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   770
    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
   771
    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
   772
    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
   773
      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
   774
    have **: "Y i \<in> B i \<and> Y i \<subseteq> X i" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   775
      apply (cases "i \<in> I")
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   776
      unfolding Y_def using * that apply (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   777
      apply (metis (no_types, lifting) someI, metis (no_types, lifting) someI_ex subset_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   778
      unfolding B_def apply simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   779
      unfolding I_def apply auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   780
      done
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   781
    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
   782
      unfolding Y_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   783
    then have ***: "finite {i. Y i \<noteq> UNIV}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   784
      unfolding I_def using H(3) rev_finite_subset by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   785
    have "(\<forall>i. Y i \<in> B i) \<and> finite {i. Y i \<noteq> UNIV}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   786
      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
   787
    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
   788
      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
   789
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   790
    have "Y i \<subseteq> X i" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   791
      apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   792
    then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   793
    then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   794
    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
   795
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   796
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   797
  show "\<exists>L. (\<forall>(i::nat). x \<in> L i \<and> open (L i)) \<and> (\<forall>U. open U \<and> x \<in> U \<longrightarrow> (\<exists>i. L i \<subseteq> U))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   798
    apply (rule first_countableI[of K])
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   799
    using \<open>countable K\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> x \<in> k\<close> \<open>\<And>k. k \<in> K \<Longrightarrow> open k\<close> Inc by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   800
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   801
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   802
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
   803
  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
   804
          topological_basis K \<and> countable K \<and>
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
   805
          (\<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
   806
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   807
  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
   808
    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
   809
  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
   810
  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
   811
  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
   812
    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
   813
  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
   814
    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
   815
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   816
  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
   817
  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
   818
    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
   819
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   820
  have "countable {X. (\<forall>(i::'a). X i \<in> B2) \<and> finite {i. X i \<noteq> UNIV}}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   821
    apply (rule countable_product_event_const) using \<open>countable B2\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   822
  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
   823
    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
   824
  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
   825
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   826
  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
   827
  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
   828
    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
   829
    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
   830
      unfolding open_fun_def by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   831
    with product_topology_open_contains_basis[OF this \<open>x \<in> U\<close>]
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   832
    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
   833
      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
   834
    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
   835
                           "\<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
   836
                           "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
   837
                           "(\<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
   838
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   839
    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
   840
    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
   841
    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
   842
    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
   843
      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
   844
    have **: "Y i \<in> B2 \<and> Y i \<subseteq> X i \<and> x i \<in> Y i" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   845
      using someI_ex[OF *]
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   846
      apply (cases "i \<in> I")
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   847
      unfolding Y_def using * apply (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   848
      unfolding B2_def I_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   849
    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
   850
      unfolding Y_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   851
    then have ***: "finite {i. Y i \<noteq> UNIV}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   852
      unfolding I_def using H(3) rev_finite_subset by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   853
    have "(\<forall>i. Y i \<in> B2) \<and> finite {i. Y i \<noteq> UNIV}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   854
      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
   855
    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
   856
      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
   857
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   858
    have "Y i \<subseteq> X i" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   859
      apply (cases "i \<in> I") using ** apply simp unfolding Y_def I_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   860
    then have "Pi\<^sub>E UNIV Y \<subseteq> Pi\<^sub>E UNIV X" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   861
    then have "Pi\<^sub>E UNIV Y \<subseteq> U" using H(4) by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   862
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   863
    have "x \<in> Pi\<^sub>E UNIV Y"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   864
      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
   865
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   866
    show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   867
      using \<open>Pi\<^sub>E UNIV Y \<in> K\<close> \<open>Pi\<^sub>E UNIV Y \<subseteq> U\<close> \<open>x \<in> Pi\<^sub>E UNIV Y\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   868
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   869
    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
   870
    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
   871
      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
   872
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   873
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   874
  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
   875
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   876
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   877
instance "fun" :: (countable, second_countable_topology) second_countable_topology
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   878
  apply standard
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   879
  using product_topology_countable_basis topological_basis_imp_subbasis by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   880
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   881
subsection \<open>Metrics on 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
   882
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   883
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   884
text \<open>In general, the product topology is not metrizable, unless the index set is countable.
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   885
When the index set is countable, essentially any (convergent) combination of the metrics on the
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   886
factors will do. We use below the simplest one, based on \<open>L\<^sup>1\<close>, but \<open>L\<^sup>2\<close> would also work,
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   887
for instance.
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   888
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   889
What is not completely trivial is that the distance thus defined induces the same topology
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   890
as the product topology. This is what we have to prove to show that we have an instance
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   891
of \<^class>\<open>metric_space\<close>.
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   892
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   893
The proofs below would work verbatim for general countable products of metric spaces. However,
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   894
since distances are only implemented in terms of type classes, we only develop the theory
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   895
for countable products of the same space.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   896
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   897
instantiation "fun" :: (countable, metric_space) metric_space
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   898
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   899
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   900
definition\<^marker>\<open>tag important\<close> dist_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
   901
  "dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   902
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70086
diff changeset
   903
definition\<^marker>\<open>tag important\<close> uniformity_fun_def:
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69030
diff changeset
   904
  "(uniformity::(('a \<Rightarrow> 'b) \<times> ('a \<Rightarrow> 'b)) filter) = (INF e\<in>{0<..}. principal {(x, y). dist (x::('a\<Rightarrow>'b)) y < e})"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   905
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   906
text \<open>Except for the first one, the auxiliary lemmas below are only useful when proving the
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   907
instance: once it is proved, they become trivial consequences of the general theory of metric
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   908
spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   909
to do this.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   910
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   911
lemma dist_fun_le_dist_first_terms:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   912
  "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   913
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   914
  have "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   915
          = (\<Sum>n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   916
    by (rule suminf_cong, simp add: power_add)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   917
  also have "... = (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   918
    apply (rule suminf_mult)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   919
    by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   920
  also have "... \<le> (1/2)^(Suc N) * (\<Sum>n. (1 / 2) ^ n)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   921
    apply (simp, rule suminf_le, simp)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   922
    by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   923
  also have "... = (1/2)^(Suc N) * 2"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   924
    using suminf_geometric[of "1/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
   925
  also have "... = (1/2)^N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   926
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   927
  finally have *: "(\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) \<le> (1/2)^N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   928
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   929
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   930
  define M where "M = Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   931
  have "dist (x (from_nat 0)) (y (from_nat 0)) \<le> M"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   932
    unfolding M_def by (rule Max_ge, auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   933
  then have [simp]: "M \<ge> 0" by (meson dual_order.trans zero_le_dist)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   934
  have "dist (x (from_nat n)) (y (from_nat n)) \<le> M" if "n\<le>N" for n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   935
    unfolding M_def apply (rule Max_ge) using that by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   936
  then have i: "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le> M" if "n\<le>N" for n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   937
    using that by force
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   938
  have "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le>
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   939
      (\<Sum>n< Suc N. M * (1 / 2) ^ n)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   940
    by (rule sum_mono, simp add: i)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   941
  also have "... = M * (\<Sum>n<Suc N. (1 / 2) ^ n)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   942
    by (rule sum_distrib_left[symmetric])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   943
  also have "... \<le> M * (\<Sum>n. (1 / 2) ^ n)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   944
    by (rule mult_left_mono, rule sum_le_suminf, auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   945
  also have "... = M * 2"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   946
    using suminf_geometric[of "1/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
   947
  finally have **: "(\<Sum>n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \<le> 2 * M"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   948
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   949
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   950
  have "dist x y = (\<Sum>n. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   951
    unfolding dist_fun_def by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   952
  also have "... = (\<Sum>n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   953
                  + (\<Sum>n<Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   954
    apply (rule suminf_split_initial_segment)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   955
    by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   956
  also have "... \<le> 2 * M + (1/2)^N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   957
    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
   958
  finally show ?thesis unfolding M_def by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   959
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   960
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   961
lemma open_fun_contains_ball_aux:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   962
  assumes "open (U::(('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
   963
          "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
   964
  shows "\<exists>e>0. {y. dist x y < e} \<subseteq> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   965
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   966
  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
   967
    using open_fun_def 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
   968
  obtain X where H: "Pi\<^sub>E UNIV X \<subseteq> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   969
                    "\<And>i. openin euclidean (X i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   970
                    "finite {i. X i \<noteq> topspace euclidean}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   971
                    "x \<in> Pi\<^sub>E UNIV X"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   972
    using product_topology_open_contains_basis[OF * \<open>x \<in> U\<close>] by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   973
  define I where "I = {i. X i \<noteq> topspace euclidean}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   974
  have "finite I" unfolding I_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
   975
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   976
    fix i
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   977
    have "x i \<in> X i" using \<open>x \<in> U\<close> H 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
   978
    then have "\<exists>e. e>0 \<and> ball (x i) e \<subseteq> X i"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   979
      using \<open>openin euclidean (X i)\<close> open_openin open_contains_ball 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
   980
    then obtain e where "e>0" "ball (x i) e \<subseteq> X i" by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   981
    define f where "f = min e (1/2)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   982
    have "f>0" "f<1" unfolding f_def using \<open>e>0\<close> by auto
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   983
    moreover have "ball (x i) f \<subseteq> X i" unfolding f_def using \<open>ball (x i) e \<subseteq> X 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
   984
    ultimately have "\<exists>f. f > 0 \<and> f < 1 \<and> ball (x i) f \<subseteq> X 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
   985
  } note * = this
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   986
  have "\<exists>e. \<forall>i. e i > 0 \<and> e i < 1 \<and> ball (x i) (e i) \<subseteq> X i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   987
    by (rule choice, auto simp add: *)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   988
  then obtain e where "\<And>i. e i > 0" "\<And>i. e i < 1" "\<And>i. ball (x i) (e i) \<subseteq> X i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   989
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   990
  define m where "m = Min {(1/2)^(to_nat i) * e i|i. 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
   991
  have "m > 0" if "I\<noteq>{}"
65036
ab7e11730ad8 Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
   992
    unfolding m_def Min_gr_iff using \<open>finite I\<close> \<open>I \<noteq> {}\<close> \<open>\<And>i. e i > 0\<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
   993
  moreover have "{y. dist x y < m} \<subseteq> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   994
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   995
    fix y assume "dist x y < m"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   996
    have "y i \<in> X i" if "i \<in> 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
   997
    proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   998
      have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   999
        by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1000
      define n where "n = to_nat i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1001
      have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 < m"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1002
        using \<open>dist x y < m\<close> unfolding dist_fun_def using sum_le_suminf[OF *, of "{n}"] 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
  1003
      then have "(1/2)^(to_nat i) * min (dist (x i) (y i)) 1 < m"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1004
        using \<open>n = to_nat 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
  1005
      also have "... \<le> (1/2)^(to_nat i) * e i"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1006
        unfolding m_def apply (rule Min_le) using \<open>finite I\<close> \<open>i \<in> 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
  1007
      ultimately have "min (dist (x i) (y i)) 1 < e i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1008
        by (auto simp add: divide_simps)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1009
      then have "dist (x i) (y i) < e i"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1010
        using \<open>e i < 1\<close> by auto
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1011
      then show "y i \<in> X i" using \<open>ball (x i) (e i) \<subseteq> X 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
  1012
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1013
    then have "y \<in> Pi\<^sub>E UNIV X" using H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1014
    then show "y \<in> U" using \<open>Pi\<^sub>E UNIV X \<subseteq> 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
  1015
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1016
  ultimately have *: "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I \<noteq> {}" using that by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1017
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1018
  have "U = UNIV" if "I = {}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1019
    using that H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1020
  then have "\<exists>m>0. {y. dist x y < m} \<subseteq> U" if "I = {}" using that zero_less_one by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1021
  then show "\<exists>m>0. {y. dist x y < m} \<subseteq> U" using * by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1022
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1023
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1024
lemma ball_fun_contains_open_aux:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1025
  fixes x::"('a \<Rightarrow> 'b)" and e::real
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1026
  assumes "e>0"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1027
  shows "\<exists>U. open U \<and> x \<in> U \<and> U \<subseteq> {y. dist x y < e}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1028
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1029
  have "\<exists>N::nat. 2^N > 8/e"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1030
    by (simp add: real_arch_pow)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1031
  then obtain N::nat where "2^N > 8/e" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1032
  define f where "f = e/4"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1033
  have [simp]: "e>0" "f > 0" unfolding f_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
  1034
  define X::"('a \<Rightarrow> 'b set)" where "X = (\<lambda>i. if (\<exists>n\<le>N. i = from_nat n) then {z. dist (x i) z < f} else UNIV)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1035
  have "{i. X i \<noteq> UNIV} \<subseteq> from_nat`{0..N}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1036
    unfolding X_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
  1037
  then have "finite {i. X i \<noteq> topspace euclidean}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1038
    unfolding topspace_euclidean 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
  1039
  have "\<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
  1040
    unfolding X_def using metric_space_class.open_ball open_UNIV by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1041
  then have "\<And>i. openin euclidean (X i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1042
    using open_openin by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1043
  define U where "U = Pi\<^sub>E UNIV X"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1044
  have "open U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1045
    unfolding open_fun_def product_topology_def apply (rule topology_generated_by_Basis)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1046
    unfolding U_def using \<open>\<And>i. openin euclidean (X i)\<close> \<open>finite {i. X i \<noteq> topspace euclidean}\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1047
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1048
  moreover have "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
  1049
    unfolding U_def X_def by (auto simp add: PiE_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1050
  moreover have "dist x y < e" if "y \<in> U" for y
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1051
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1052
    have *: "dist (x (from_nat n)) (y (from_nat n)) \<le> f" if "n \<le> N" for n
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1053
      using \<open>y \<in> U\<close> unfolding U_def apply (auto simp add: 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
  1054
      unfolding X_def using that by (metis less_imp_le mem_Collect_eq)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1055
    have **: "Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} \<le> f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1056
      apply (rule Max.boundedI) 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
  1057
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1058
    have "dist x y \<le> 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \<le> N} + (1/2)^N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1059
      by (rule dist_fun_le_dist_first_terms)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1060
    also have "... \<le> 2 * f + e / 8"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1061
      apply (rule add_mono) using ** \<open>2^N > 8/e\<close> by(auto simp add: algebra_simps divide_simps)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1062
    also have "... \<le> e/2 + e/8"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1063
      unfolding f_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
  1064
    also have "... < e"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1065
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1066
    finally show "dist x y < e" by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1067
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1068
  ultimately show ?thesis by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1069
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1070
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1071
lemma fun_open_ball_aux:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1072
  fixes U::"('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
  1073
  shows "open U \<longleftrightarrow> (\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1074
proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1075
  assume "open U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1076
  fix x assume "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
  1077
  then show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1078
    using open_fun_contains_ball_aux[OF \<open>open U\<close> \<open>x \<in> U\<close>] by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1079
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1080
  assume H: "\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> y \<in> U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1081
  define K where "K = {V. open V \<and> V \<subseteq> U}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1082
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1083
    fix x assume "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
  1084
    then obtain e where e: "e>0" "{y. dist x y < e} \<subseteq> U" using H by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1085
    then obtain V where V: "open V" "x \<in> V" "V \<subseteq> {y. dist x y < e}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1086
      using ball_fun_contains_open_aux[OF \<open>e>0\<close>, of x] 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
  1087
    have "V \<in> K"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1088
      unfolding K_def using e(2) V(1) V(3) by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1089
    then have "x \<in> (\<Union>K)" using \<open>x \<in> V\<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
  1090
  }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1091
  then have "(\<Union>K) = U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1092
    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
  1093
  moreover have "open (\<Union>K)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1094
    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
  1095
  ultimately show "open U" by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1096
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1097
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1098
instance proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1099
  fix x y::"'a \<Rightarrow> 'b" show "(dist x y = 0) = (x = y)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1100
  proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1101
    assume "x = y"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1102
    then show "dist x y = 0" unfolding dist_fun_def using \<open>x = y\<close> by auto
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1103
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1104
    assume "dist x y = 0"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1105
    have *: "summable (\<lambda>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1106
      by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1107
    have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 = 0" for n
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1108
      using \<open>dist x y = 0\<close> unfolding dist_fun_def by (simp add: "*" suminf_eq_zero_iff)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1109
    then have "dist (x (from_nat n)) (y (from_nat n)) = 0" for n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1110
      by (metis div_0 min_def nonzero_mult_div_cancel_left power_eq_0_iff
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1111
                zero_eq_1_divide_iff zero_neq_numeral)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1112
    then have "x (from_nat n) = y (from_nat n)" for n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1113
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1114
    then have "x i = y i" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1115
      by (metis from_nat_to_nat)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1116
    then show "x = y"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1117
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1118
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1119
next
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1120
  text\<open>The proof of the triangular inequality is trivial, modulo the fact that we are dealing
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1121
        with infinite series, hence we should justify the convergence at each step. In this
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1122
        respect, the following simplification rule is extremely handy.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1123
  have [simp]: "summable (\<lambda>n. (1/2)^n * min (dist (u (from_nat n)) (v (from_nat n))) 1)" for u v::"'a \<Rightarrow> 'b"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1124
    by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1125
  fix x y z::"'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
  1126
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1127
    fix n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1128
    have *: "dist (x (from_nat n)) (y (from_nat n)) \<le>
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1129
            dist (x (from_nat n)) (z (from_nat n)) + dist (y (from_nat n)) (z (from_nat n))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1130
      by (simp add: dist_triangle2)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1131
    have "0 \<le> dist (y (from_nat n)) (z (from_nat n))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1132
      using zero_le_dist by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1133
    moreover
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1134
    {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1135
      assume "min (dist (y (from_nat n)) (z (from_nat n))) 1 \<noteq> dist (y (from_nat n)) (z (from_nat n))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1136
      then have "1 \<le> min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1137
        by (metis (no_types) diff_le_eq diff_self min_def zero_le_dist zero_le_one)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1138
    }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1139
    ultimately have "min (dist (x (from_nat n)) (y (from_nat n))) 1 \<le>
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1140
            min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1141
      using * by linarith
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1142
  } note ineq = this
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1143
  have "dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1144
    unfolding dist_fun_def by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1145
  also have "... \<le> (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1146
                        + (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1147
    apply (rule suminf_le)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1148
    using ineq apply (metis (no_types, hide_lams) add.right_neutral distrib_left
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1149
      le_divide_eq_numeral1(1) mult_2_right mult_left_mono zero_le_one zero_le_power)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1150
    by (auto simp add: summable_add)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1151
  also have "... = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1152
                  + (\<Sum>n. (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1153
    by (rule suminf_add[symmetric], auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1154
  also have "... = dist x z + dist y z"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1155
    unfolding dist_fun_def by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1156
  finally show "dist x y \<le> dist x z + dist y z"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1157
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1158
next
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1159
  text\<open>Finally, we show that the topology generated by the distance and the product
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1160
        topology coincide. This is essentially contained in Lemma \<open>fun_open_ball_aux\<close>,
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1161
        except that the condition to prove is expressed with filters. To deal with this,
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1162
        we copy some mumbo jumbo from Lemma \<open>eventually_uniformity_metric\<close> in
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1163
        \<^file>\<open>~~/src/HOL/Real_Vector_Spaces.thy\<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
  1164
  fix U::"('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
  1165
  have "eventually P uniformity \<longleftrightarrow> (\<exists>e>0. \<forall>x (y::('a \<Rightarrow> 'b)). dist x y < e \<longrightarrow> P (x, y))" for P
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1166
  unfolding uniformity_fun_def apply (subst eventually_INF_base)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1167
    by (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1168
  then show "open U = (\<forall>x\<in>U. \<forall>\<^sub>F (x', y) in uniformity. x' = x \<longrightarrow> y \<in> U)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1169
    unfolding fun_open_ball_aux by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1170
qed (simp add: uniformity_fun_def)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1171
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1172
end
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1173
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1174
text \<open>Nice properties of spaces are preserved under countable products. In addition
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1175
to first countability, second countability and metrizability, as we have seen above,
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1176
completeness is also preserved, and therefore being Polish.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1177
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1178
instance "fun" :: (countable, complete_space) complete_space
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1179
proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1180
  fix u::"nat \<Rightarrow> ('a \<Rightarrow> 'b)" assume "Cauchy u"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1181
  have "Cauchy (\<lambda>n. u n 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
  1182
  unfolding cauchy_def proof (intro impI allI)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1183
    fix e assume "e>(0::real)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1184
    obtain k where "i = from_nat k"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1185
      using from_nat_to_nat[of i] by metis
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1186
    have "(1/2)^k * min e 1 > 0" using \<open>e>0\<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
  1187
    then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1188
      using \<open>Cauchy u\<close> unfolding cauchy_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
  1189
    then obtain N::nat where C: "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m) (u n) < (1/2)^k * min e 1"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1190
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1191
    have "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1192
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1193
      fix m n::nat assume "m \<ge> N" "n \<ge> N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1194
      have "(1/2)^k * min (dist (u m i) (u n i)) 1
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1195
              = sum (\<lambda>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1) {k}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1196
        using \<open>i = from_nat 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
  1197
      also have "... \<le> (\<Sum>p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1198
        apply (rule sum_le_suminf)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1199
        by (rule summable_comparison_test'[of "\<lambda>n. (1/2)^n"], auto simp add: summable_geometric_iff)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1200
      also have "... = dist (u m) (u n)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1201
        unfolding dist_fun_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
  1202
      also have "... < (1/2)^k * min e 1"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1203
        using C \<open>m \<ge> N\<close> \<open>n \<ge> N\<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
  1204
      finally have "min (dist (u m i) (u n i)) 1 < min e 1"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1205
        by (auto simp add: algebra_simps divide_simps)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1206
      then show "dist (u m i) (u n i) < e" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1207
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1208
    then show "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (u m i) (u n i) < e"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1209
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1210
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1211
  then have "\<exists>x. (\<lambda>n. u n i) \<longlonglongrightarrow> x" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1212
    using Cauchy_convergent convergent_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
  1213
  then have "\<exists>x. \<forall>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1214
    using choice by force
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1215
  then obtain x where *: "\<And>i. (\<lambda>n. u n i) \<longlonglongrightarrow> x i" by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1216
  have "u \<longlonglongrightarrow> x"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1217
  proof (rule metric_LIMSEQ_I)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1218
    fix e assume [simp]: "e>(0::real)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1219
    have i: "\<exists>K. \<forall>n\<ge>K. dist (u n i) (x i) < e/4" for i
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1220
      by (rule metric_LIMSEQ_D, auto simp add: *)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1221
    have "\<exists>K. \<forall>i. \<forall>n\<ge>K i. dist (u n i) (x i) < e/4"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1222
      apply (rule choice) using 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
  1223
    then obtain K where K: "\<And>i n. n \<ge> K i \<Longrightarrow> dist (u n i) (x i) < e/4"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1224
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1225
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1226
    have "\<exists>N::nat. 2^N > 4/e"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1227
      by (simp add: real_arch_pow)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1228
    then obtain N::nat where "2^N > 4/e" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1229
    define L where "L = Max {K (from_nat n)|n. n \<le> N}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1230
    have "dist (u k) x < e" if "k \<ge> L" for k
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1231
    proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1232
      have *: "dist (u k (from_nat n)) (x (from_nat n)) \<le> e / 4" if "n \<le> N" for n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1233
      proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1234
        have "K (from_nat n) \<le> L"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1235
          unfolding L_def apply (rule Max_ge) using \<open>n \<le> N\<close> by auto
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1236
        then have "k \<ge> K (from_nat n)" using \<open>k \<ge> L\<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
  1237
        then show ?thesis using K less_imp_le by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1238
      qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1239
      have **: "Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} \<le> e/4"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1240
        apply (rule Max.boundedI) 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
  1241
      have "dist (u k) x \<le> 2 * Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \<le> N} + (1/2)^N"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1242
        using dist_fun_le_dist_first_terms by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1243
      also have "... \<le> 2 * e/4 + e/4"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1244
        apply (rule add_mono)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1245
        using ** \<open>2^N > 4/e\<close> less_imp_le by (auto simp add: algebra_simps divide_simps)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1246
      also have "... < e" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1247
      finally 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
  1248
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1249
    then show "\<exists>L. \<forall>k\<ge>L. dist (u k) x < e" by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1250
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1251
  then show "convergent u" using convergent_def by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1252
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1253
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1254
instance "fun" :: (countable, polish_space) polish_space
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
  1255
  by standard
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1256
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1257
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1258
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
  1259
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1260
theorem Alexander_subbase:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1261
  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
  1262
    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
  1263
  shows "compact_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1264
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1265
  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
  1266
    by (simp flip: X)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1267
  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
  1268
    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
  1269
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1270
    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
  1271
    have 1: "\<A> \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1272
      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
  1273
    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
  1274
      unfolding \<A>_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1275
    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
  1276
      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
  1277
        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
  1278
      have "\<C> \<subseteq> \<A>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1279
        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
  1280
      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
  1281
        by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1282
      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
  1283
      proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1284
        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
  1285
          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
  1286
        then show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1287
          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
  1288
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1289
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1290
    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
  1291
      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
  1292
    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
  1293
                        \<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
  1294
           \<Longrightarrow> \<W> = \<K>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1295
      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
  1296
      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
  1297
      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
  1298
    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
  1299
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1300
      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
  1301
        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
  1302
      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
  1303
        by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1304
      then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1305
        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
  1306
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1307
    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
  1308
      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
  1309
    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
  1310
      by (metis openin_subset)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1311
    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
  1312
      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
  1313
    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
  1314
      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
  1315
    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
  1316
      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
  1317
      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
  1318
    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
  1319
      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
  1320
    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
  1321
      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
  1322
    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
  1323
      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
  1324
    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
  1325
      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
  1326
    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
  1327
      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
  1328
    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
  1329
    proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1330
      fix b
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1331
      assume "b \<in> \<B>'"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1332
      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
  1333
      proof (rule *)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1334
        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
  1335
          using that
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1336
        proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1337
          have "b \<in> \<B>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1338
            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
  1339
          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
  1340
            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
  1341
          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
  1342
            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
  1343
          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
  1344
            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
  1345
            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
  1346
            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
  1347
            done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1348
          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
  1349
            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
  1350
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1351
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1352
        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
  1353
          using top by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1354
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1355
        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
  1356
        proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1357
          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
  1358
            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
  1359
          then show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1360
            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
  1361
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1362
      qed auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1363
      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
  1364
        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
  1365
        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
  1366
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1367
    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
  1368
      by metis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1369
    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
  1370
    show False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1371
    proof (rule non)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1372
      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
  1373
        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
  1374
      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
  1375
        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
  1376
      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
  1377
      show "?\<D> \<subseteq> \<K>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1378
        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
  1379
      show "finite ?\<D>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1380
        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
  1381
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1382
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1383
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1384
    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
  1385
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1386
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1387
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1388
corollary Alexander_subbase_alt:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1389
  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
  1390
  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
  1391
   and X: "topology
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1392
              (arbitrary union_of
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1393
                   (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
  1394
 shows "compact_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1395
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1396
  have "topspace X = U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1397
    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
  1398
  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
  1399
    unfolding relative_to_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1400
    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
  1401
  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
  1402
    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
  1403
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1404
    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
  1405
      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
  1406
    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
  1407
      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
  1408
    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
  1409
      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
  1410
    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
  1411
      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
  1412
    then show ?thesis
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1413
      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
  1414
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1415
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1416
    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
  1417
     apply (simp flip: X)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1418
     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
  1419
    apply (blast intro: *)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1420
    done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1421
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1422
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1423
proposition continuous_map_componentwise:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1424
   "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
  1425
    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
  1426
    (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
  1427
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
  1428
  case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1429
  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
  1430
    by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1431
  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
  1432
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1433
    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
  1434
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1435
      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
  1436
        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
  1437
        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
  1438
        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
  1439
        done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1440
      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
  1441
        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
  1442
      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
  1443
        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
  1444
      ultimately show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1445
        by metis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1446
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1447
    with that
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1448
    show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1449
      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
  1450
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1451
  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
  1452
  proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1453
    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
  1454
      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
  1455
    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
  1456
      by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1457
    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
  1458
      by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1459
    show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1460
      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
  1461
    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
  1462
      fix \<T>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1463
      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
  1464
                  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
  1465
      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
  1466
      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
  1467
        fix S T
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1468
        assume "T \<in> \<T>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1469
        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
  1470
          "\<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
  1471
          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
  1472
        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
  1473
          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
  1474
          unfolding 3
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1475
          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
  1476
          done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1477
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1478
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1479
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1480
  ultimately show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1481
    by metis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1482
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1483
  case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1484
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1485
    by (auto simp: continuous_map_def PiE_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1486
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1487
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1488
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1489
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
  1490
   "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
  1491
  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
  1492
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1493
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
  1494
   "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
  1495
  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
  1496
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1497
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
  1498
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1499
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
  1500
  assumes "i \<in> I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1501
  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
  1502
  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
  1503
proof clarify
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1504
  fix \<V>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1505
  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
  1506
               (finite intersection_of
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1507
                (\<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
  1508
                           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
  1509
  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
  1510
  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
  1511
    fix S V
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1512
    assume "V \<in> \<V>"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1513
    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
  1514
      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
  1515
      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
  1516
      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
  1517
      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
  1518
    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
  1519
    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
  1520
      fix x f
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1521
      assume "f \<in> V"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1522
      let ?T = "{a \<in> topspace(Y i).
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1523
                   (\<lambda>j. if j = i then a
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1524
                        else if j \<in> I then f j else undefined) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1525
      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
  1526
      proof (intro exI conjI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1527
        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
  1528
        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
  1529
          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
  1530
          proof (cases "k=i")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1531
            case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1532
            then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1533
              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
  1534
          next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1535
            case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1536
            then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1537
              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
  1538
          qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1539
          then show "continuous_map (Y i) (product_topology Y I)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1540
                  (\<lambda>x j. if j = i then x else if j \<in> I then f j else undefined)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1541
            by (auto simp: continuous_map_componentwise assms extensional_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1542
        next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1543
          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
  1544
            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
  1545
          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
  1546
                         if "\<F> \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1547
          proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1548
            show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1549
            proof (rule openin_Inter)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1550
              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
  1551
                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
  1552
                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
  1553
                apply (rename_tac F)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1554
                apply (rule_tac x=F in exI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1555
                using subsetD [OF \<F>]
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1556
                apply (force intro: finite_intersection_of_inc)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1557
                done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1558
            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
  1559
          qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1560
          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
  1561
            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
  1562
        qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1563
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1564
        have eqf: "(\<lambda>j. if j = i then f i else if j \<in> I then f j else undefined) = f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1565
          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
  1566
        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
  1567
          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
  1568
      next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1569
        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
  1570
          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
  1571
      qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1572
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1573
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1574
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1575
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1576
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
  1577
  assumes "i \<in> I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1578
  shows "(retraction_map (product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1579
         (topspace (product_topology X I) = {}) \<longrightarrow> topspace (X i) = {})"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1580
  (is "?lhs = ?rhs")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1581
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1582
  assume ?lhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1583
  then show ?rhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1584
    using retraction_imp_surjective_map by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1585
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1586
  assume R: ?rhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1587
  show ?lhs
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1588
  proof (cases "topspace (product_topology X I) = {}")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1589
    case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1590
    then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1591
      using R by (auto simp: retraction_map_def retraction_maps_def continuous_map_on_empty)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1592
  next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1593
    case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1594
    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
  1595
      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
  1596
    proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1597
      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
  1598
        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
  1599
      show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1600
        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
  1601
        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
  1602
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1603
    show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1604
      using \<open>i \<in> I\<close> False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1605
      by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1606
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1607
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
  1608
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
  1609
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
  1610
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1611
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
  1612
  "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
  1613
        PiE I S = {} \<or>
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1614
        finite {i \<in> I. ~(S i = topspace(X i))} \<and> (\<forall>i \<in> I. openin (X i) (S i))"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1615
  (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
  1616
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
  1617
  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
  1618
  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
  1619
  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
  1620
    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
  1621
    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
  1622
    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
  1623
      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
  1624
    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
  1625
      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
  1626
      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
  1627
      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
  1628
    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
  1629
      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
  1630
    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
  1631
    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
  1632
      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
  1633
        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
  1634
        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
  1635
        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
  1636
    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
  1637
      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
  1638
      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
  1639
      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
  1640
        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
  1641
        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
  1642
        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
  1643
        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
  1644
        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
  1645
    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
  1646
  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
  1647
    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
  1648
    then 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
  1649
      apply (simp only: openin_product_topology)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1650
      apply (rule arbitrary_union_of_inc)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1651
      apply (auto simp: product_topology_base_alt)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1652
      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
  1653
  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
  1654
  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
  1655
    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
  1656
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
  1657
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1658
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1659
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
  1660
   "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
  1661
  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
  1662
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1663
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
  1664
   "compact_space(product_topology X I) \<longleftrightarrow>
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1665
    topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. compact_space(X i))"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1666
    (is "?lhs = ?rhs")
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1667
proof (cases "topspace(product_topology X I) = {}")
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1668
  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
  1669
  then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1670
    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
  1671
  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
  1672
  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
  1673
    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
  1674
    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
  1675
    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
  1676
      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
  1677
      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
  1678
      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
  1679
        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
  1680
      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
  1681
      have "\<And>x. x \<in> topspace (X i) \<Longrightarrow> x \<in> (\<lambda>f. f i) ` (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1682
        using \<open>i \<in> I\<close> z
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1683
        apply (rule_tac x="\<lambda>j. if j = i then x else if j \<in> I then z j else undefined" in image_eqI, auto)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1684
        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
  1685
      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
  1686
        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
  1687
      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
  1688
        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
  1689
    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
  1690
  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
  1691
    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
  1692
    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
  1693
    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
  1694
      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
  1695
      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
  1696
        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
  1697
    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
  1698
      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
  1699
      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
  1700
        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
  1701
      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
  1702
        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
  1703
      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
  1704
        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
  1705
        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
  1706
        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
  1707
        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
  1708
          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
  1709
            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
  1710
        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
  1711
          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
  1712
          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
  1713
          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
  1714
          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
  1715
          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
  1716
            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
  1717
            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
  1718
                   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
  1719
              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
  1720
            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
  1721
                            \<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
  1722
              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
  1723
            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
  1724
              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
  1725
            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
  1726
              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
  1727
            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
  1728
              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
  1729
              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
  1730
          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
  1731
            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
  1732
            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
  1733
              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
  1734
            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
  1735
              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
  1736
            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
  1737
              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
  1738
            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
  1739
              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
  1740
            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
  1741
              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
  1742
          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
  1743
        qed (simp add: product_topology)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1744
      qed (simp add: compact_space_topspace_empty)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1745
    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
  1746
  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
  1747
qed (simp add: compact_space_topspace_empty)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1748
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1749
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
  1750
   "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
  1751
        PiE I S = {} \<or> (\<forall>i \<in> I. compactin (X i) (S i))"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1752
  by (auto simp: compactin_subspace subtopology_PiE subset_PiE compact_space_product_topology
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1753
       PiE_eq_empty_iff)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1754
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1755
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
  1756
   "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
  1757
        \<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
  1758
  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
  1759
  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
  1760
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1761
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
  1762
   "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
  1763
  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
  1764
  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
  1765
  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
  1766
     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
  1767
  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
  1768
   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
  1769
  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
  1770
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1771
subsection\<open>Relationship with connected spaces, paths, etc.\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1772
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1773
proposition connected_space_product_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1774
   "connected_space(product_topology X I) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1775
    (\<Pi>\<^sub>E i\<in>I. topspace (X i)) = {} \<or> (\<forall>i \<in> I. connected_space(X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1776
  (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1777
proof (cases ?eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1778
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1779
  moreover have "?lhs = ?rhs"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1780
  proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1781
    assume ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1782
    moreover
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1783
    have "connectedin(X i) (topspace(X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1784
      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
  1785
    proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1786
      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
  1787
        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
  1788
      show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1789
        using connectedin_continuous_map_image [OF cm ci] \<open>i \<in> I\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1790
        by (simp add: False image_projection_PiE)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1791
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1792
    ultimately show ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1793
      by (meson connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1794
  next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1795
    assume cs [rule_format]: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1796
    have False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1797
      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
  1798
        and U: "openin (product_topology X I) U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1799
        and V: "openin (product_topology X I) V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1800
        and "U \<noteq> {}" "V \<noteq> {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1801
      for U V
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1802
    proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1803
      obtain f where "f \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1804
        using \<open>U \<noteq> {}\<close> by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1805
      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
  1806
        using U openin_subset by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1807
      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
  1808
        using U V openin_subset by blast+
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1809
      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
  1810
      proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1811
        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
  1812
            (\<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
  1813
          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
  1814
        then obtain \<T> where "finite \<T>"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1815
          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
  1816
          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
  1817
          and ftop: "f \<in> topspace (product_topology X I)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1818
          and fint: "f \<in> \<Inter> \<T>"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1819
          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
  1820
        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
  1821
        obtain L where "finite L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1822
          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
  1823
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1824
          show "finite ?L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1825
          proof (rule finite_Union)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1826
            fix M
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1827
            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
  1828
            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
  1829
              by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1830
            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
  1831
              using t by meson
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1832
            then have "f j \<in> V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1833
              using \<open>C \<in> \<T>\<close> fint by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1834
            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
  1835
              using that
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1836
              apply (clarsimp simp add: set_eq_iff)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1837
              apply (rule_tac x="\<lambda>m. if m = k then x else f m" in image_eqI, auto)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1838
              done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1839
            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
  1840
              using Ceq by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1841
            then show "finite M"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1842
              using C finite_subset by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1843
          qed (use \<open>finite \<T>\<close> in blast)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1844
        next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1845
          fix i U
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1846
          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
  1847
          then show "i \<in> ?L"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1848
            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
  1849
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1850
        show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1851
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1852
          fix h
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1853
          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
  1854
          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
  1855
          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
  1856
            unfolding g_def using f h by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1857
          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
  1858
            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
  1859
          ultimately have "g \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1860
            using subU by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1861
          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
  1862
            using that
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1863
          proof (induction arbitrary: h)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1864
            case empty
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1865
            then show ?case
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1866
              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
  1867
          next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1868
            case (insert i M)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1869
            define f where "f \<equiv> \<lambda>j. if j = i then g i else h j"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1870
            have fin: "f \<in> PiE I (topspace \<circ> X)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1871
              unfolding f_def using gin insert.prems(1) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1872
            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
  1873
              unfolding f_def using insert.prems(2) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1874
            have "f \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1875
              using insert.IH [OF fin subM] .
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1876
            show ?case
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1877
            proof (cases "h \<in> V")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1878
              case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1879
              show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1880
              proof (cases "i \<in> I")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1881
                case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1882
                let ?U = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> U}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1883
                let ?V = "{x \<in> topspace(X i). (\<lambda>j. if j = i then x else h j) \<in> V}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1884
                have False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1885
                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
  1886
                  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
  1887
                    using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1888
                  then have cm: "continuous_map (X i) (product_topology X I) (\<lambda>x j. if j = i then x else h j)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1889
                    using \<open>i \<in> I\<close> insert.prems(1)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1890
                    by (auto simp: continuous_map_componentwise extensional_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1891
                  show "openin (X i) ?U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1892
                    by (rule openin_continuous_map_preimage [OF cm U])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1893
                  show "openin (X i) ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1894
                    by (rule openin_continuous_map_preimage [OF cm V])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1895
                  show "topspace (X i) \<subseteq> ?U \<union> ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1896
                  proof clarsimp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1897
                    fix x
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1898
                    assume "x \<in> topspace (X i)" and "(\<lambda>j. if j = i then x else h j) \<notin> V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1899
                    with True subUV \<open>h \<in> Pi\<^sub>E I (topspace \<circ> X)\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1900
                    show "(\<lambda>j. if j = i then x else h j) \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1901
                      by (drule_tac c="(\<lambda>j. if j = i then x else h j)" in subsetD) auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1902
                  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1903
                  show "?U \<inter> ?V = {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1904
                    using disj by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1905
                  show "?U \<noteq> {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1906
                    using \<open>U \<noteq> {}\<close> f_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1907
                  proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1908
                    have "(\<lambda>j. if j = i then g i else h j) \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1909
                      using \<open>f \<in> U\<close> f_def by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1910
                    moreover have "f i \<in> topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1911
                      by (metis PiE_iff True comp_apply fin)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1912
                    ultimately have "\<exists>b. b \<in> topspace (X i) \<and> (\<lambda>a. if a = i then b else h a) \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1913
                      using f_def by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1914
                    then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1915
                      by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1916
                  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1917
                  have "(\<lambda>j. if j = i then h i else h j) = h"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1918
                    by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1919
                  moreover have "h i \<in> topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1920
                    using True insert.prems(1) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1921
                  ultimately show "?V \<noteq> {}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1922
                    using \<open>h \<in> V\<close>  by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1923
                qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1924
                then show ?thesis ..
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1925
              next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1926
                case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1927
                show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1928
                proof (cases "h = f")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1929
                  case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1930
                  show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1931
                    by (rule insert.IH [OF insert.prems(1)]) (simp add: True subM)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1932
                next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1933
                  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1934
                  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1935
                    using gin insert.prems(1) \<open>i \<notin> I\<close> unfolding f_def by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1936
                qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1937
              qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1938
            next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1939
              case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1940
              then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1941
                using subUV insert.prems(1) by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1942
            qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1943
          qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1944
          then show "h \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1945
            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
  1946
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1947
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1948
      ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1949
        using disj inf_absorb2 \<open>V \<noteq> {}\<close> by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1950
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1951
    then show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1952
      unfolding connected_space_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1953
      by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1954
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1955
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1956
    by simp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1957
qed (simp add: connected_space_topspace_empty)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1958
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1959
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1960
lemma connectedin_PiE:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1961
   "connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1962
        PiE I S = {} \<or> (\<forall>i \<in> I. connectedin (X i) (S i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1963
  by (fastforce simp add: connectedin_def subtopology_PiE connected_space_product_topology subset_PiE PiE_eq_empty_iff
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1964
      topspace_subtopology_subset)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1965
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1966
lemma path_connected_space_product_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1967
   "path_connected_space(product_topology X I) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1968
    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
  1969
 (is "?lhs \<longleftrightarrow> ?eq \<or> ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1970
proof (cases ?eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1971
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1972
  moreover have "?lhs = ?rhs"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1973
  proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1974
    assume L: ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1975
    show ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1976
    proof (clarsimp simp flip: path_connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1977
      fix i :: "'a"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1978
      assume "i \<in> I"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1979
      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
  1980
        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
  1981
      show "path_connectedin (X i) (topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1982
        using path_connectedin_continuous_map_image [OF cm L [unfolded path_connectedin_topspace [symmetric]]]
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1983
        by (metis \<open>i \<in> I\<close> False retraction_imp_surjective_map retraction_map_product_projection)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1984
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1985
  next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1986
    assume R [rule_format]: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1987
    show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1988
      unfolding path_connected_space_def topspace_product_topology
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1989
    proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1990
      fix x y
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1991
      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
  1992
      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
  1993
        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
  1994
      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
  1995
        by metis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1996
      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
  1997
        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
  1998
        apply (force simp: pathin_def continuous_map_componentwise)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1999
        done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2000
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2001
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2002
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2003
    by simp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2004
qed (simp add: path_connected_space_topspace_empty)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2005
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2006
lemma path_connectedin_PiE:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2007
   "path_connectedin (product_topology X I) (PiE I S) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2008
    PiE I S = {} \<or> (\<forall>i \<in> I. path_connectedin (X i) (S i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2009
  by (fastforce simp add: path_connectedin_def subtopology_PiE path_connected_space_product_topology subset_PiE PiE_eq_empty_iff topspace_subtopology_subset)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2010
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2011
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
  2012
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2013
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
  2014
  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
  2015
  shows "quotient_map(product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2016
           (topspace(product_topology X I) = {} \<longrightarrow> topspace(X i) = {})"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2017
         (is "?lhs = ?rhs")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2018
proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2019
  assume ?lhs with assms show ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2020
    by (auto simp: continuous_open_quotient_map open_map_product_projection)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2021
next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2022
  assume ?rhs with assms show ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2023
    by (auto simp: Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2024
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2025
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2026
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
  2027
  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
  2028
  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
  2029
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2030
  have "quotient_map (product_topology X I) (X i) (\<lambda>x. x i)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2031
    using assms by (force simp add: quotient_map_product_projection PiE_eq_empty_iff)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2032
  moreover
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2033
  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
  2034
    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
  2035
  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
  2036
    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
  2037
    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
  2038
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2039
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2040
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
  2041
  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
  2042
    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
  2043
                      \<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
  2044
               (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
  2045
    and PQ:  "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2046
  shows "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. Q(X i))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2047
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2048
  have "Q(X i)" if "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<noteq> {}" "i \<in> I" for i
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2049
  proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2050
    from that obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" 
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2051
      by force
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2052
    have "?SX f i homeomorphic_space X i"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2053
      apply (simp add: subtopology_PiE )
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2054
      using product_topology_homeomorphic_component [OF \<open>i \<in> I\<close>, of "\<lambda>j. subtopology (X j) (if j = i then topspace (X i) else {f j})"]
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2055
      using f by fastforce
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2056
    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
  2057
      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
  2058
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69945
diff changeset
  2059
  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
  2060
qed
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  2061
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  2062
end