src/HOL/Analysis/Function_Topology.thy
author paulson <lp15@cam.ac.uk>
Thu, 07 Mar 2019 14:08:05 +0000
changeset 69874 11065b70407d
parent 69745 aec42cee2521
child 69918 eddcc7c726f3
permissions -rw-r--r--
new material for Analysis
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
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
     6
imports Topology_Euclidean_Space Bounded_Linear_Function Finite_Product_Measure 
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
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
    10
section \<open>Product Topology\<close> (* FIX  see comments by the author *)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    11
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    12
text \<open>We want to define the 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]
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    45
\<open>continuous_on_topo T1 T2 f =
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 =
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    51
    continuous_on_topo (subtopology euclidean 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
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    53
In fact, I need \<open>continuous_on_topo\<close> to express the continuity of the projection on subfactors
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
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    55
the above equivalence in Lemma~\<open>continuous_on_continuous_on_topo\<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>Topology without 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
    65
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
    66
subsubsection%important \<open>The topology generated by some (open) subsets\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    67
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
    68
text \<open>In the definition below of a generated topology, the \<open>Empty\<close> case is not necessary,
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    69
as it follows from \<open>UN\<close> taking for \<open>K\<close> the empty set. However, it is convenient to have,
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    70
and is never a problem in proofs, so I prefer to write it down explicitly.
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    71
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    72
We do not require \<open>UNIV\<close> to be an open set, as this will not be the case in applications. (We are
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    73
thinking of a topology on a subset of \<open>UNIV\<close>, the remaining part of \<open>UNIV\<close> being irrelevant.)\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    74
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    75
inductive generate_topology_on for S where
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    76
  Empty: "generate_topology_on S {}"
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    77
| Int: "generate_topology_on S a \<Longrightarrow> generate_topology_on S b \<Longrightarrow> generate_topology_on S (a \<inter> b)"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    78
| UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology_on S k) \<Longrightarrow> generate_topology_on S (\<Union>K)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    79
| Basis: "s \<in> S \<Longrightarrow> generate_topology_on S s"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    80
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    81
lemma istopology_generate_topology_on:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    82
  "istopology (generate_topology_on S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    83
unfolding istopology_def by (auto intro: generate_topology_on.intros)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    84
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    85
text \<open>The basic property of the topology generated by a set \<open>S\<close> is that it is the
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
    86
smallest topology containing all the elements of \<open>S\<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
    87
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    88
lemma generate_topology_on_coarsest:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    89
  assumes "istopology T"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    90
          "\<And>s. s \<in> S \<Longrightarrow> T s"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    91
          "generate_topology_on S s0"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    92
  shows "T s0"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    93
using assms(3) apply (induct rule: generate_topology_on.induct)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    94
using assms(1) assms(2) unfolding istopology_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
    95
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    96
abbreviation%unimportant topology_generated_by::"('a set set) \<Rightarrow> ('a topology)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
    97
  where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
    98
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    99
lemma openin_topology_generated_by_iff:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   100
  "openin (topology_generated_by S) s \<longleftrightarrow> generate_topology_on S s"
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   101
  using topology_inverse'[OF istopology_generate_topology_on[of S]] 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
   102
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   103
lemma openin_topology_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
   104
  "openin (topology_generated_by S) s \<Longrightarrow> generate_topology_on S s"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   105
using openin_topology_generated_by_iff by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   106
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   107
lemma 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
   108
  "topspace (topology_generated_by S) = (\<Union>S)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   109
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   110
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   111
    fix s assume "openin (topology_generated_by S) s"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   112
    then have "generate_topology_on S s" by (rule openin_topology_generated_by)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   113
    then have "s \<subseteq> (\<Union>S)" by (induct, auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   114
  }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   115
  then show "topspace (topology_generated_by S) \<subseteq> (\<Union>S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   116
    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
   117
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   118
  have "generate_topology_on S (\<Union>S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   119
    using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   120
  then show "(\<Union>S) \<subseteq> topspace (topology_generated_by S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   121
    unfolding topspace_def using openin_topology_generated_by_iff by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   122
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   123
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   124
lemma topology_generated_by_Basis:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   125
  "s \<in> S \<Longrightarrow> openin (topology_generated_by S) s"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   126
  by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   127
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   128
lemma generate_topology_on_Inter:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   129
  "\<lbrakk>finite \<F>; \<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on \<S> K; \<F> \<noteq> {}\<rbrakk> \<Longrightarrow> generate_topology_on \<S> (\<Inter>\<F>)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   130
  by (induction \<F> rule: finite_induct; force intro: generate_topology_on.intros)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   131
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   132
subsection\<open>Topology bases and sub-bases\<close>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   133
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   134
lemma istopology_base_alt:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   135
   "istopology (arbitrary union_of P) \<longleftrightarrow>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   136
    (\<forall>S T. (arbitrary union_of P) S \<and> (arbitrary union_of P) T
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   137
           \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   138
  by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   139
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   140
lemma istopology_base_eq:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   141
   "istopology (arbitrary union_of P) \<longleftrightarrow>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   142
    (\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   143
  by (simp add: istopology_base_alt arbitrary_union_of_Int_eq)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   144
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   145
lemma istopology_base:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   146
   "(\<And>S T. \<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<inter> T)) \<Longrightarrow> istopology (arbitrary union_of P)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   147
  by (simp add: arbitrary_def istopology_base_eq union_of_inc)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   148
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   149
lemma openin_topology_base_unique:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   150
   "openin X = arbitrary union_of P \<longleftrightarrow>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   151
        (\<forall>V. P V \<longrightarrow> openin X V) \<and> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V. P V \<and> x \<in> V \<and> V \<subseteq> U))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   152
   (is "?lhs = ?rhs")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   153
proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   154
  assume ?lhs
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   155
  then show ?rhs
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   156
    by (auto simp: union_of_def arbitrary_def)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   157
next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   158
  assume R: ?rhs
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   159
  then have *: "\<exists>\<U>\<subseteq>Collect P. \<Union>\<U> = S" if "openin X S" for S
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   160
    using that by (rule_tac x="{V. P V \<and> V \<subseteq> S}" in exI) fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   161
  from R show ?lhs
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   162
    by (fastforce simp add: union_of_def arbitrary_def intro: *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   163
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   164
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   165
lemma topology_base_unique:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   166
   "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S;
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   167
     \<And>U x. \<lbrakk>openin X U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>B. P B \<and> x \<in> B \<and> B \<subseteq> U\<rbrakk>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   168
    \<Longrightarrow> topology(arbitrary union_of P) = X"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   169
  by (metis openin_topology_base_unique openin_inverse [of X])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   170
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   171
lemma topology_bases_eq_aux:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   172
   "\<lbrakk>(arbitrary union_of P) S;
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   173
     \<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U\<rbrakk>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   174
        \<Longrightarrow> (arbitrary union_of Q) S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   175
  by (metis arbitrary_union_of_alt arbitrary_union_of_idempot)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   176
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   177
lemma topology_bases_eq:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   178
   "\<lbrakk>\<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U;
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   179
    \<And>V x. \<lbrakk>Q V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> V\<rbrakk>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   180
        \<Longrightarrow> topology (arbitrary union_of P) =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   181
            topology (arbitrary union_of Q)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   182
  by (fastforce intro:  arg_cong [where f=topology]  elim: topology_bases_eq_aux)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   183
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   184
lemma istopology_subbase:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   185
   "istopology (arbitrary union_of (finite intersection_of P relative_to S))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   186
  by (simp add: finite_intersection_of_Int istopology_base relative_to_Int)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   187
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   188
lemma openin_subbase:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   189
  "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   190
   \<longleftrightarrow> (arbitrary union_of (finite intersection_of B relative_to U)) S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   191
  by (simp add: istopology_subbase topology_inverse')
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   192
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   193
lemma topspace_subbase [simp]:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   194
   "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   195
proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   196
  show "?lhs \<subseteq> U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   197
    by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   198
  show "U \<subseteq> ?lhs"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   199
    by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase 
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   200
              openin_subset relative_to_inc subset_UNIV topology_inverse')
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   201
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   202
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   203
lemma minimal_topology_subbase:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   204
   "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S; openin X U;
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   205
     openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\<rbrakk>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   206
    \<Longrightarrow> openin X S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   207
  apply (simp add: istopology_subbase topology_inverse)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   208
  apply (simp add: union_of_def intersection_of_def relative_to_def)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   209
  apply (blast intro: openin_Int_Inter)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   210
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   211
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   212
lemma istopology_subbase_UNIV:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   213
   "istopology (arbitrary union_of (finite intersection_of P))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   214
  by (simp add: istopology_base finite_intersection_of_Int)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   215
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   216
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   217
lemma generate_topology_on_eq:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   218
  "generate_topology_on S = arbitrary union_of finite' intersection_of (\<lambda>x. x \<in> S)" (is "?lhs = ?rhs")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   219
proof (intro ext iffI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   220
  fix A
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   221
  assume "?lhs A"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   222
  then show "?rhs A"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   223
  proof induction
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   224
    case (Int a b)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   225
    then show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   226
      by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   227
  next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   228
    case (UN K)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   229
    then show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   230
      by (simp add: arbitrary_union_of_Union)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   231
  next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   232
    case (Basis s)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   233
    then show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   234
      by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   235
  qed auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   236
next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   237
  fix A
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   238
  assume "?rhs A"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   239
  then obtain \<U> where \<U>: "\<And>T. T \<in> \<U> \<Longrightarrow> \<exists>\<F>. finite' \<F> \<and> \<F> \<subseteq> S \<and> \<Inter>\<F> = T" and eq: "A = \<Union>\<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   240
    unfolding union_of_def intersection_of_def by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   241
  show "?lhs A"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   242
    unfolding eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   243
  proof (rule generate_topology_on.UN)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   244
    fix T
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   245
    assume "T \<in> \<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   246
    with \<U> obtain \<F> where "finite' \<F>" "\<F> \<subseteq> S" "\<Inter>\<F> = T"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   247
      by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   248
    have "generate_topology_on S (\<Inter>\<F>)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   249
    proof (rule generate_topology_on_Inter)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   250
      show "finite \<F>" "\<F> \<noteq> {}"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   251
        by (auto simp: \<open>finite' \<F>\<close>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   252
      show "\<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on S K"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   253
        by (metis \<open>\<F> \<subseteq> S\<close> generate_topology_on.simps subset_iff)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   254
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   255
    then show "generate_topology_on S T"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   256
      using \<open>\<Inter>\<F> = T\<close> by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   257
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   258
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   259
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
   260
subsubsection%important \<open>Continuity\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   261
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   262
text \<open>We will need to deal with continuous maps in terms of topologies and not in terms
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   263
of type classes, as defined below.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   264
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
   265
definition%important continuous_on_topo::"'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   266
  where "continuous_on_topo T1 T2 f = ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<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
   267
                                      \<and> (f`(topspace T1) \<subseteq> (topspace T2)))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   268
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   269
lemma continuous_on_continuous_on_topo:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   270
  "continuous_on s f \<longleftrightarrow> continuous_on_topo (subtopology euclidean s) euclidean f"
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
   271
  by (auto simp: continuous_on_topo_def Int_commute continuous_openin_preimage_eq)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   272
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   273
lemma continuous_on_topo_UNIV:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   274
  "continuous_on UNIV f \<longleftrightarrow> continuous_on_topo euclidean euclidean f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   275
using continuous_on_continuous_on_topo[of UNIV f] subtopology_UNIV[of euclidean] by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   276
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   277
lemma continuous_on_topo_open [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   278
  "continuous_on_topo T1 T2 f \<Longrightarrow> openin T2 U \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   279
  unfolding continuous_on_topo_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
   280
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   281
lemma continuous_on_topo_topspace [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   282
  "continuous_on_topo T1 T2 f \<Longrightarrow> f`(topspace T1) \<subseteq> (topspace T2)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   283
unfolding continuous_on_topo_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
   284
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   285
lemma continuous_on_generated_topo_iff:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   286
  "continuous_on_topo T1 (topology_generated_by S) f \<longleftrightarrow>
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   287
      ((\<forall>U. U \<in> S \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) \<and> (f`(topspace T1) \<subseteq> (\<Union> S)))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   288
unfolding continuous_on_topo_def topology_generated_by_topspace
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   289
proof (auto simp add: topology_generated_by_Basis)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   290
  assume H: "\<forall>U. U \<in> S \<longrightarrow> openin T1 (f -` U \<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
   291
  fix U assume "openin (topology_generated_by S) U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   292
  then have "generate_topology_on S U" by (rule openin_topology_generated_by)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   293
  then show "openin T1 (f -` U \<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
   294
  proof (induct)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   295
    fix a b
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   296
    assume H: "openin T1 (f -` a \<inter> topspace T1)" "openin T1 (f -` b \<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
   297
    have "f -` (a \<inter> b) \<inter> topspace T1 = (f-`a \<inter> topspace T1) \<inter> (f-`b \<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
   298
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   299
    then show "openin T1 (f -` (a \<inter> b) \<inter> topspace T1)" 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
   300
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   301
    fix K
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   302
    assume H: "openin T1 (f -` k \<inter> topspace T1)" 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
   303
    define L where "L = {f -` k \<inter> topspace T1|k. k \<in> K}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   304
    have *: "openin T1 l" if "l \<in>L" for l using that H unfolding L_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
   305
    have "openin T1 (\<Union>L)" using openin_Union[OF *] by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   306
    moreover have "(\<Union>L) = (f -` \<Union>K \<inter> topspace T1)" unfolding L_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
   307
    ultimately show "openin T1 (f -` \<Union>K \<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
   308
  qed (auto simp add: H)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   309
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   310
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   311
lemma 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
   312
  assumes "\<And>U. U \<in>S \<Longrightarrow> openin T1 (f-`U \<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
   313
          "f`(topspace T1) \<subseteq> (\<Union> S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   314
  shows "continuous_on_topo T1 (topology_generated_by S) f"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   315
  using assms continuous_on_generated_topo_iff 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
   316
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   317
proposition continuous_on_topo_compose:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   318
  assumes "continuous_on_topo T1 T2 f" "continuous_on_topo T2 T3 g"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   319
  shows "continuous_on_topo T1 T3 (g o f)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   320
  using assms unfolding continuous_on_topo_def
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   321
proof (auto)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   322
  fix U :: "'c set"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   323
  assume H: "openin T3 U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   324
  have "openin T1 (f -` (g -` U \<inter> topspace T2) \<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
   325
    using H assms by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   326
  moreover have "f -` (g -` U \<inter> topspace T2) \<inter> topspace T1 = (g \<circ> f) -` U \<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
   327
    using H assms continuous_on_topo_topspace by fastforce
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   328
  ultimately show "openin T1 ((g \<circ> f) -` U \<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
   329
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   330
qed (blast)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   331
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   332
lemma continuous_on_topo_preimage_topspace [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   333
  assumes "continuous_on_topo T1 T2 f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   334
  shows "f-`(topspace T2) \<inter> topspace T1 = topspace T1"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   335
using assms unfolding continuous_on_topo_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
   336
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   337
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
   338
subsubsection%important \<open>Pullback 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
   339
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   340
text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> is
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   341
a special case of this notion, pulling back by the identity. We introduce the general notion as
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   342
we will need it to define the strong operator topology on the space of continuous linear operators,
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   343
by pulling back the product topology on the space of all functions.\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   344
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   345
text \<open>\<open>pullback_topology A f T\<close> is the pullback of the topology \<open>T\<close> by the map \<open>f\<close> on
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
   346
the set \<open>A\<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
   347
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
   348
definition%important pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<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
   349
  where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   350
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   351
lemma istopology_pullback_topology:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   352
  "istopology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   353
  unfolding istopology_def proof (auto)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   354
  fix K assume "\<forall>S\<in>K. \<exists>U. openin T U \<and> S = f -` U \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   355
  then have "\<exists>U. \<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   356
    by (rule bchoice)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   357
  then obtain U where U: "\<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   358
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   359
  define V where "V = (\<Union>S\<in>K. U S)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   360
  have "openin T V" "\<Union>K = f -` V \<inter> A" unfolding V_def using U by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   361
  then show "\<exists>V. openin T V \<and> \<Union>K = f -` V \<inter> A" by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   362
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   363
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   364
lemma openin_pullback_topology:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   365
  "openin (pullback_topology A f T) S \<longleftrightarrow> (\<exists>U. openin T U \<and> S = f-`U \<inter> A)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   366
unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   367
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   368
lemma topspace_pullback_topology:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   369
  "topspace (pullback_topology A f T) = f-`(topspace T) \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   370
by (auto simp add: topspace_def openin_pullback_topology)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   371
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   372
proposition continuous_on_topo_pullback [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   373
  assumes "continuous_on_topo T1 T2 g"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   374
  shows "continuous_on_topo (pullback_topology A f T1) T2 (g o f)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   375
unfolding continuous_on_topo_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   376
proof (auto)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   377
  fix U::"'b set" assume "openin T2 U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   378
  then have "openin T1 (g-`U \<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
   379
    using assms unfolding continuous_on_topo_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
   380
  have "(g o f)-`U \<inter> topspace (pullback_topology A f T1) = (g o f)-`U \<inter> A \<inter> f-`(topspace T1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   381
    unfolding topspace_pullback_topology by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   382
  also have "... = f-`(g-`U \<inter> topspace T1) \<inter> A "
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   383
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   384
  also have "openin (pullback_topology A f T1) (...)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   385
    unfolding openin_pullback_topology using \<open>openin T1 (g-`U \<inter> topspace T1)\<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
   386
  finally show "openin (pullback_topology A f T1) ((g \<circ> f) -` U \<inter> topspace (pullback_topology A f T1))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   387
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   388
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   389
  fix x assume "x \<in> topspace (pullback_topology A f T1)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   390
  then have "f 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
   391
    unfolding topspace_pullback_topology by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   392
  then show "g (f x) \<in> topspace T2"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   393
    using assms unfolding continuous_on_topo_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
   394
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   395
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   396
proposition continuous_on_topo_pullback' [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   397
  assumes "continuous_on_topo T1 T2 (f o g)" "topspace T1 \<subseteq> g-`A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   398
  shows "continuous_on_topo T1 (pullback_topology A f T2) g"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   399
unfolding continuous_on_topo_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   400
proof (auto)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   401
  fix U assume "openin (pullback_topology A f T2) U"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   402
  then have "\<exists>V. openin T2 V \<and> U = f-`V \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   403
    unfolding openin_pullback_topology by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   404
  then obtain V where "openin T2 V" "U = f-`V \<inter> A"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   405
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   406
  then have "g -` U \<inter> topspace T1 = g-`(f-`V \<inter> A) \<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
   407
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   408
  also have "... = (f o g)-`V \<inter> (g-`A \<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
   409
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   410
  also have "... = (f o g)-`V \<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
   411
    using 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
   412
  also have "openin T1 (...)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   413
    using assms(1) \<open>openin T2 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
   414
  finally show "openin T1 (g -` 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
   415
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   416
  fix x assume "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
   417
  have "(f o g) x \<in> topspace T2"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   418
    using assms(1) \<open>x \<in> topspace T1\<close> unfolding continuous_on_topo_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
   419
  then have "g x \<in> f-`(topspace T2)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   420
    unfolding comp_def by blast
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   421
  moreover have "g x \<in> A" using assms(2) \<open>x \<in> topspace T1\<close> 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
   422
  ultimately show "g x \<in> topspace (pullback_topology A f T2)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   423
    unfolding topspace_pullback_topology by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   424
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   425
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   426
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   427
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
   428
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   429
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
   430
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
   431
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
   432
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
   433
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
   434
(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
   435
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
   436
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   437
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
   438
\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   439
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
   440
definition%important 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
   441
  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
   442
    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
   443
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   444
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
   445
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
   446
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   447
proposition product_topology:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   448
  "product_topology X I =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   449
   topology
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   450
    (arbitrary union_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   451
       ((finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   452
        (\<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
   453
        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
   454
  (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
   455
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   456
  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
   457
  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
   458
  proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   459
    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
   460
      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
   461
    proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   462
      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
   463
        using \<U> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   464
      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
   465
        by metis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   466
      obtain U where "U \<in> \<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   467
        using \<open>\<U> \<noteq> {}\<close> by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   468
      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
   469
      show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   470
      proof (intro conjI exI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   471
        show "finite (\<Union>U\<in>\<U>. ?F U)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   472
          using Y \<open>finite' \<U>\<close> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   473
        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
   474
        proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   475
          have *: "f \<in> U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   476
            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
   477
              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
   478
          proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   479
            have "Pi\<^sub>E I (Y U) = U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   480
              using Y \<open>U \<in> \<U>\<close> by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   481
            then show "f \<in> U"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   482
              using that unfolding PiE_def Pi_def by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   483
          qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   484
          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
   485
            by (auto simp: PiE_iff *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   486
          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
   487
            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
   488
        qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   489
      qed (use Y openin_subset in \<open>blast+\<close>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   490
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   491
    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
   492
      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
   493
    proof (cases "\<U>={}")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   494
      case True
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   495
      then show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   496
        apply (rule_tac x="{?TOP}" in exI, simp)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   497
        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
   498
        apply force
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   499
        done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   500
    next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   501
      case False
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   502
      then obtain U where "U \<in> \<U>"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   503
        by blast
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   504
      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
   505
        using \<U> by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   506
      then obtain J Y where
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   507
        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
   508
        by metis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   509
      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
   510
      show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   511
      proof (intro conjI exI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   512
        show "finite (?G ` \<U>)" "?G ` \<U> \<noteq> {}"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   513
          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
   514
        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
   515
          using Y by force
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   516
        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
   517
          apply clarsimp
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   518
          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
   519
          apply (auto simp: *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   520
          done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   521
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   522
        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
   523
        proof
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   524
          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
   525
            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
   526
            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
   527
          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
   528
            using \<open>U \<in> \<U>\<close>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   529
            by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   530
          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
   531
            using PiE_mem Y by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   532
          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
   533
            by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   534
        qed (use Y in fastforce)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   535
      qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   536
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   537
    show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   538
      unfolding relative_to_def intersection_of_def
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   539
      by (safe; blast dest!: 1 2)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   540
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   541
  show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   542
    unfolding product_topology_def generate_topology_on_eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   543
    apply (rule arg_cong [where f = topology])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   544
    apply (rule arg_cong [where f = "(union_of)arbitrary"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   545
    apply (force simp: *)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   546
    done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   547
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   548
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   549
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
   550
  "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
   551
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   552
  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
   553
    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
   554
    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
   555
  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
   556
    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
   557
  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
   558
    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
   559
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   560
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   561
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
   562
  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
   563
  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
   564
  unfolding product_topology_def
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   565
  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
   566
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   567
proposition product_topology_open_contains_basis:
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   568
  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
   569
  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
   570
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   571
  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
   572
    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
   573
  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
   574
  proof induction
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   575
    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
   576
    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
   577
      "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
   578
      "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
   579
      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
   580
    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
   581
    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
   582
      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
   583
    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
   584
    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
   585
      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
   586
    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
   587
      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
   588
      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
   589
    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
   590
      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
   591
    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
   592
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   593
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   594
    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
   595
    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
   596
    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
   597
      by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   598
    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
   599
      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
    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
   601
      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
   602
    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
   603
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   604
  qed auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   605
  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
   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
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   608
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   609
lemma openin_product_topology:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   610
   "openin (product_topology X I) =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   611
    arbitrary union_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   612
          ((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
   613
           relative_to topspace (product_topology X I))"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   614
  apply (subst product_topology)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   615
  apply (simp add: topspace_product_topology topology_inverse' [OF istopology_subbase])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   616
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   617
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   618
lemma subtopology_PiE:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   619
  "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
   620
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   621
  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
   622
  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
   623
  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
   624
        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
   625
    by (rule finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   626
  also have "\<dots> = finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   627
                      ((\<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
   628
                       relative_to ?X \<inter> Pi\<^sub>E I S)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   629
                   relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   630
    apply (rule arg_cong2 [where f = "(relative_to)"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   631
     apply (rule arg_cong [where f = "(intersection_of)finite"])
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   632
     apply (rule ext)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   633
     apply (auto simp: relative_to_def intersection_of_def)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   634
    done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   635
  finally
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   636
  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
   637
        finite intersection_of
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   638
          (\<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
   639
          relative_to ?X \<inter> Pi\<^sub>E I S"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   640
    by (metis finite_intersection_of_relative_to)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   641
  then show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   642
  unfolding topology_eq
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   643
  apply clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   644
  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
   645
  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
   646
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   647
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   648
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   649
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   650
lemma product_topology_base_alt:
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   651
   "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
   652
    relative_to topspace(product_topology X I) =
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   653
    (\<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
   654
   (is "?lhs = ?rhs")
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   655
proof -
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   656
  have "(\<forall>F. ?lhs F \<longrightarrow> ?rhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   657
    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
   658
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   659
    fix \<F>
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   660
    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
   661
    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
   662
          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
   663
    proof (induction)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   664
      case (insert F \<F>)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   665
      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
   666
        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
   667
        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
   668
        by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   669
      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
   670
        using insert by auto
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   671
      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
   672
      show ?case
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   673
      proof (intro exI conjI)
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69722
diff changeset
   674
        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
   675
        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
   676
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   677
        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
   678
          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
   679
      next
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   680
        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
   681
          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
   682
      qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   683
    qed (auto intro: dest: not_finite_existsD)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   684
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   685
  moreover have "(\<forall>F. ?rhs F \<longrightarrow> ?lhs F)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   686
  proof clarify
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   687
    fix U :: "'a \<Rightarrow> 'b set"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   688
    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
   689
    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
   690
    show "?lhs (Pi\<^sub>E I U)"
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   691
      unfolding relative_to_def topspace_product_topology
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   692
    proof (intro exI conjI)
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   693
      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
   694
        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
   695
      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
   696
        using ope openin_subset by fastforce
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   697
    qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   698
  qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   699
  ultimately show ?thesis
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   700
    by meson
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   701
qed
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   702
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   703
lemma closure_of_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   704
  "(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
   705
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   706
  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
   707
       \<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
   708
    (is "?lhs = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   709
    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
   710
  proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   711
    assume L[rule_format]: ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   712
    show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   713
    proof clarify
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   714
      fix i T
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   715
      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
   716
      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
   717
        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
   718
      then show "False"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   719
        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
   720
        by (auto simp: top ext PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   721
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   722
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   723
    assume R [rule_format]: ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   724
    show ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   725
    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
   726
      fix \<U> U
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   727
      assume
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   728
        \<U>: "\<U> \<subseteq> Collect
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   729
           (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
   730
            (\<Pi>\<^sub>E i\<in>I. topspace (X i)))" and
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   731
        "f \<in> U" "U \<in> \<U>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   732
      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
   733
                 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
   734
        by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   735
      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
   736
      obtain \<T> where "finite \<T>"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   737
        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
   738
        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
   739
        apply (clarsimp simp add: relative_to_def intersection_of_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   740
        apply (rule that, auto dest!: subsetD)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   741
        done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   742
      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
   743
        by (auto simp: PiE_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   744
      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
   745
             \<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
   746
        if "i \<in> I" for i
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   747
      proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   748
        have "finite ((\<lambda>U. {x. x i \<in> U}) -` \<T>)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   749
        proof (rule finite_vimageI [OF \<open>finite \<T>\<close>])
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   750
          show "inj (\<lambda>U. {x. x i \<in> U})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   751
            by (auto simp: inj_on_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   752
        qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   753
        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
   754
          by (rule rev_finite_subset) auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   755
        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
   756
          by (rule openin_Inter) (auto simp: fin)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   757
        then show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   758
          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
   759
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   760
      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
   761
      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
   762
        using R [OF _ *] unfolding \<Phi>_def by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   763
      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
   764
        by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   765
      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
   766
      proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   767
        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
   768
        proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   769
          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
   770
            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
   771
          then show "restrict \<theta> I \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   772
            using subU by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   773
        qed (rule \<open>U \<in> \<U>\<close>)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   774
      next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   775
        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
   776
          using \<theta> by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   777
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   778
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   779
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   780
  show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   781
    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
   782
    by metis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   783
qed
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   784
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   785
corollary closedin_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   786
   "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
   787
  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
   788
  apply (metis closure_of_empty)
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   789
  done
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   790
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   791
corollary closedin_product_topology_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   792
   "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
   793
  using PiE_singleton closedin_product_topology [of X I]
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   794
  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
   795
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   796
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
   797
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
   798
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   799
lemma continuous_on_topo_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
   800
  assumes "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
   801
  shows "continuous_on_topo (product_topology T I) (T i) (\<lambda>x. x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   802
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   803
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   804
    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
   805
    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
   806
    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
   807
      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
   808
      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
   809
    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
   810
      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
   811
    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
   812
      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
   813
      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
   814
      apply (subst *)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   815
      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
   816
  }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   817
  then show ?thesis unfolding continuous_on_topo_def
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   818
    by (auto simp add: assms topspace_product_topology 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
   819
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   820
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   821
lemma continuous_on_topo_coordinatewise_then_product [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   822
  assumes "\<And>i. i \<in> I \<Longrightarrow> continuous_on_topo T1 (T i) (\<lambda>x. f x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   823
          "\<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
   824
  shows "continuous_on_topo T1 (product_topology T I) f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   825
unfolding product_topology_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   826
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
   827
  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
   828
  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
   829
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   830
  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
   831
  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
   832
  have "(\<lambda>x. f x i)-`(topspace(T i)) \<inter> topspace T1 = topspace T1" 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
   833
    using that assms(1) by (simp add: continuous_on_topo_preimage_topspace)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   834
  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
   835
    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
   836
  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
   837
    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
   838
  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
   839
    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
   840
  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
   841
    apply (rule openin_INT)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   842
    apply (simp add: \<open>finite J\<close>)
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   843
    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
   844
  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
   845
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   846
  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
   847
    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
   848
    apply (subst product_topology_def[symmetric])
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   849
    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
   850
    apply (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
   851
    using assms unfolding continuous_on_topo_def by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   852
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   853
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   854
lemma continuous_on_topo_product_then_coordinatewise [intro]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   855
  assumes "continuous_on_topo T1 (product_topology T I) f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   856
  shows "\<And>i. i \<in> I \<Longrightarrow> continuous_on_topo T1 (T i) (\<lambda>x. f x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   857
        "\<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
   858
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   859
  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
   860
  have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f" 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
  also have "continuous_on_topo T1 (T i) (...)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   862
    apply (rule continuous_on_topo_compose[of _ "product_topology T I"])
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   863
    using assms \<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
   864
  ultimately show "continuous_on_topo T1 (T i) (\<lambda>x. f x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   865
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   866
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   867
  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
   868
  have "f x \<in> topspace (product_topology T I)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   869
    using assms \<open>x \<in> topspace T1\<close> unfolding continuous_on_topo_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
   870
  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
   871
    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
   872
  then show "f x i = undefined"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   873
    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
   874
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   875
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   876
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
   877
  assumes "J \<subseteq> I"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   878
  shows "continuous_on_topo (product_topology T I) (product_topology T J) (\<lambda>x. restrict x J)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   879
proof (rule continuous_on_topo_coordinatewise_then_product)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   880
  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
   881
  then have "(\<lambda>x. restrict x J i) = (\<lambda>x. x i)" 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
   882
  then show "continuous_on_topo (product_topology T I) (T i) (\<lambda>x. restrict x J i)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   883
    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
   884
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   885
  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
   886
  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
   887
    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
   888
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   889
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   890
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
   891
subsubsection%important \<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
   892
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   893
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
   894
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   895
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
   896
definition%important 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
   897
  "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
   898
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   899
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
   900
  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
   901
    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
   902
  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
   903
    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
   904
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
   905
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   906
end
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   907
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   908
lemma 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
   909
  "euclidean = product_topology (\<lambda>i. euclidean::('b::topological_space) topology) UNIV"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   910
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
   911
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   912
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
   913
  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
   914
  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
   915
  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
   916
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   917
  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
   918
  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
   919
  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
   920
  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
   921
    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
   922
  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
   923
    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
   924
  have "open (Pi\<^sub>E UNIV X)"
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
   925
    unfolding open_fun_def 
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
   926
    by (simp_all add: * ** 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
   927
  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
   928
    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
   929
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   930
      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
   931
      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
   932
      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
   933
      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
   934
        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
   935
      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
   936
        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
   937
      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
   938
        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
   939
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   940
  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
   941
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   942
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   943
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
   944
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
   945
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   946
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
   947
  "continuous_on UNIV (\<lambda>x. x i::('b::topological_space))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   948
unfolding continuous_on_topo_UNIV euclidean_product_topology
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   949
by (rule continuous_on_topo_product_coordinates, simp)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   950
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   951
lemma continuous_on_coordinatewise_then_product [intro, continuous_intros]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   952
  assumes "\<And>i. continuous_on UNIV (\<lambda>x. f x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   953
  shows "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
   954
using assms unfolding continuous_on_topo_UNIV euclidean_product_topology
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 continuous_on_topo_coordinatewise_then_product, simp)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   956
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   957
lemma continuous_on_product_then_coordinatewise:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   958
  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
   959
  shows "continuous_on UNIV (\<lambda>x. f x i)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   960
using assms unfolding continuous_on_topo_UNIV euclidean_product_topology
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   961
by (rule continuous_on_topo_product_then_coordinatewise(1), simp)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   962
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   963
lemma continuous_on_product_coordinatewise_iff:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   964
  "continuous_on UNIV f \<longleftrightarrow> (\<forall>i. continuous_on UNIV (\<lambda>x. f x i))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   965
by (auto intro: continuous_on_product_then_coordinatewise)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   966
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69710
diff changeset
   967
subsubsection%important \<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
   968
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
   969
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
   970
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
   971
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
   972
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   973
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
   974
  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
   975
  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
   976
  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
   977
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   978
  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
   979
                  \<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
   980
    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
   981
  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
   982
  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
   983
    case 0
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   984
    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
   985
      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
   986
    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
   987
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   988
    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
   989
    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
   990
      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
   991
    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
   992
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
   993
      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
   994
      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
   995
      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
   996
        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
   997
      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
   998
        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
   999
      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
  1000
        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
  1001
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1002
    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
  1003
      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
  1004
    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
  1005
      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
  1006
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1007
  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
  1008
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1009
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1010
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
  1011
  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
  1012
  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
  1013
  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
  1014
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1015
  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
  1016
  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
  1017
  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
  1018
  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
  1019
  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
  1020
        \<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
  1021
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1022
    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
  1023
    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
  1024
    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
  1025
    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
  1026
    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
  1027
      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
  1028
    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
  1029
      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
  1030
    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
  1031
      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
    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
  1033
      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
  1034
    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
  1035
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1036
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1037
  then show ?thesis
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1038
    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
  1039
    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
  1040
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1041
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1042
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
  1043
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1044
  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
  1045
  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
  1046
    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
  1047
  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
  1048
                                                  "\<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
  1049
                                                  "\<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
  1050
    by metis
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1051
  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
  1052
  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
  1053
  have "countable (B i)" for i unfolding 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
  1054
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1055
  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
  1056
  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
  1057
    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
  1058
  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
  1059
  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
  1060
    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
  1061
  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
  1062
    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
  1063
  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
  1064
  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
  1065
    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
  1066
  have "open 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
  1067
    using that unfolding open_fun_def K_def B_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
  1068
    apply (rule product_topology_basis)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1069
    unfolding topspace_euclidean by (auto, metis imageE open_UNIV A(2))
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1070
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1071
  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
  1072
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1073
    have "openin (product_topology (\<lambda>i. euclidean) UNIV) U" "x \<in> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1074
      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
  1075
    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
  1076
    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
  1077
      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
  1078
    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
  1079
                           "\<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
  1080
                           "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
  1081
                           "(\<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
  1082
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1083
    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
  1084
    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
  1085
    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
  1086
      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
  1087
    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
  1088
      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
  1089
      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
  1090
      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
  1091
      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
  1092
      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
  1093
      done
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1094
    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
  1095
      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
  1096
    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
  1097
      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
  1098
    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
  1099
      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
  1100
    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
  1101
      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
  1102
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1103
    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
  1104
      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
  1105
    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
  1106
    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
  1107
    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
  1108
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1109
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1110
  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
  1111
    apply (rule first_countableI[of K])
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1112
    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
  1113
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1114
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1115
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
  1116
  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
  1117
          topological_basis K \<and> countable K \<and>
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
  1118
          (\<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
  1119
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1120
  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
  1121
    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
  1122
  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
  1123
  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
  1124
  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
  1125
    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
  1126
  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
  1127
    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
  1128
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1129
  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
  1130
  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
  1131
    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
  1132
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1133
  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
  1134
    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
  1135
  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
  1136
    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
  1137
  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
  1138
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1139
  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
  1140
  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
  1141
    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
  1142
    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
  1143
      unfolding open_fun_def by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1144
    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
  1145
    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
  1146
      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
  1147
    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
  1148
                           "\<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
  1149
                           "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
  1150
                           "(\<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
  1151
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1152
    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
  1153
    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
  1154
    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
  1155
    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
  1156
      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
  1157
    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
  1158
      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
  1159
      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
  1160
      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
  1161
      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
  1162
    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
  1163
      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
  1164
    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
  1165
      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
  1166
    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
  1167
      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
  1168
    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
  1169
      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
  1170
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1171
    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
  1172
      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
  1173
    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
  1174
    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
  1175
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1176
    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
  1177
      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
  1178
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1179
    show "\<exists>V\<in>K. x \<in> V \<and> V \<subseteq> U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1180
      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
  1181
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1182
    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
  1183
    show "open U"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1184
      using \<open>U \<in> K\<close> unfolding open_fun_def K_def apply (auto)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1185
      apply (rule product_topology_basis)
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  1186
      using \<open>\<And>V. V \<in> B2 \<Longrightarrow> open V\<close> open_UNIV 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
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1188
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1189
  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
  1190
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1191
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1192
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
  1193
  apply standard
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1194
  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
  1195
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1196
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
  1197
subsection \<open>The strong operator topology on continuous linear operators\<close> (* FIX ME mv*)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1198
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1199
text \<open>Let \<open>'a\<close> and \<open>'b\<close> be two normed real vector spaces. Then the space of linear continuous
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1200
operators from \<open>'a\<close> to \<open>'b\<close> has a canonical norm, and therefore a canonical corresponding topology
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1201
(the type classes instantiation are given in \<^file>\<open>Bounded_Linear_Function.thy\<close>).
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1202
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1203
However, there is another topology on this space, the strong operator topology, where \<open>T\<^sub>n\<close> tends to
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1204
\<open>T\<close> iff, for all \<open>x\<close> in \<open>'a\<close>, then \<open>T\<^sub>n x\<close> tends to \<open>T x\<close>. This is precisely the product topology
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1205
where the target space is endowed with the norm topology. It is especially useful when \<open>'b\<close> is the set
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1206
of real numbers, since then this topology is compact.
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1207
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1208
We can not implement it using type classes as there is already a topology, but at least we
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1209
can define it as a topology.
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1210
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1211
Note that there is yet another (common and useful) topology on operator spaces, the weak operator
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1212
topology, defined analogously using the product topology, but where the target space is given the
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1213
weak-* topology, i.e., the pullback of the weak topology on the bidual of the space under the
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1214
canonical embedding of a space into its bidual. We do not define it there, although it could also be
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1215
defined analogously.
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1216
\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1217
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
  1218
definition%important strong_operator_topology::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) topology"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1219
where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1220
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1221
lemma strong_operator_topology_topspace:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1222
  "topspace strong_operator_topology = UNIV"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1223
unfolding strong_operator_topology_def topspace_pullback_topology topspace_euclidean by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1224
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1225
lemma strong_operator_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
  1226
  fixes f::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector)" and U::"'i \<Rightarrow> 'b set" and x::"'i \<Rightarrow> 'a"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1227
  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
  1228
  shows "openin strong_operator_topology {f. \<forall>i\<in>I. blinfun_apply f (x i) \<in> U i}"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1229
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1230
  have "open {g::('a\<Rightarrow>'b). \<forall>i\<in>I. g (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
  1231
    by (rule product_topology_basis'[OF assms])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1232
  moreover have "{f. \<forall>i\<in>I. blinfun_apply 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
  1233
                = blinfun_apply-`{g::('a\<Rightarrow>'b). \<forall>i\<in>I. g (x i) \<in> U i} \<inter> UNIV"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1234
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1235
  ultimately show ?thesis
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69683
diff changeset
  1236
    unfolding strong_operator_topology_def by (subst openin_pullback_topology) 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
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1238
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1239
lemma strong_operator_topology_continuous_evaluation:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1240
  "continuous_on_topo strong_operator_topology euclidean (\<lambda>f. blinfun_apply f x)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1241
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1242
  have "continuous_on_topo strong_operator_topology euclidean ((\<lambda>f. f x) o blinfun_apply)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1243
    unfolding strong_operator_topology_def apply (rule continuous_on_topo_pullback)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1244
    using continuous_on_topo_UNIV continuous_on_product_coordinates by fastforce
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1245
  then show ?thesis unfolding comp_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
  1246
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1247
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1248
lemma continuous_on_strong_operator_topo_iff_coordinatewise:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1249
  "continuous_on_topo T strong_operator_topology f
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1250
    \<longleftrightarrow> (\<forall>x. continuous_on_topo T euclidean (\<lambda>y. blinfun_apply (f y) x))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1251
proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1252
  fix x::"'b"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1253
  assume "continuous_on_topo T strong_operator_topology f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1254
  with continuous_on_topo_compose[OF this strong_operator_topology_continuous_evaluation]
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1255
  have "continuous_on_topo T euclidean ((\<lambda>z. blinfun_apply z x) o f)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1256
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1257
  then show "continuous_on_topo T euclidean (\<lambda>y. blinfun_apply (f y) x)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1258
    unfolding comp_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
  1259
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1260
  assume *: "\<forall>x. continuous_on_topo T euclidean (\<lambda>y. blinfun_apply (f y) x)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1261
  have "continuous_on_topo T euclidean (blinfun_apply o f)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1262
    unfolding euclidean_product_topology
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1263
    apply (rule continuous_on_topo_coordinatewise_then_product, auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1264
    using * unfolding comp_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
  1265
  show "continuous_on_topo T strong_operator_topology f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1266
    unfolding strong_operator_topology_def
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1267
    apply (rule continuous_on_topo_pullback')
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1268
    by (auto simp add: \<open>continuous_on_topo T euclidean (blinfun_apply o 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
  1269
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1270
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1271
lemma strong_operator_topology_weaker_than_euclidean:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1272
  "continuous_on_topo euclidean strong_operator_topology (\<lambda>f. f)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1273
  by (subst continuous_on_strong_operator_topo_iff_coordinatewise,
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1274
    auto simp add: continuous_on_topo_UNIV[symmetric] linear_continuous_on)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1275
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1276
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
  1277
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
  1278
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1279
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1280
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
  1281
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
  1282
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
  1283
for instance.
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1284
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1285
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
  1286
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
  1287
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
  1288
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1289
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
  1290
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
  1291
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
  1292
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1293
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
  1294
begin
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1295
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
  1296
definition%important 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
  1297
  "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
  1298
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 66827
diff changeset
  1299
definition%important uniformity_fun_def:
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69030
diff changeset
  1300
  "(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
  1301
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1302
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
  1303
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
  1304
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
  1305
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
  1306
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1307
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
  1308
  "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
  1309
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1310
  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
  1311
          = (\<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
  1312
    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
  1313
  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
  1314
    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
  1315
    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
  1316
  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
  1317
    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
  1318
    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
  1319
  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
  1320
    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
  1321
  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
  1322
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1323
  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
  1324
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1325
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1326
  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
  1327
  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
  1328
    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
  1329
  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
  1330
  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
  1331
    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
  1332
  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
  1333
    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
  1334
  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
  1335
      (\<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
  1336
    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
  1337
  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
  1338
    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
  1339
  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
  1340
    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
  1341
  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
  1342
    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
  1343
  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
  1344
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1345
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1346
  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
  1347
    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
  1348
  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
  1349
                  + (\<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
  1350
    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
  1351
    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
  1352
  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
  1353
    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
  1354
  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
  1355
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1356
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1357
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
  1358
  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
  1359
          "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
  1360
  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
  1361
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1362
  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
  1363
    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
  1364
  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
  1365
                    "\<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
  1366
                    "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
  1367
                    "x \<in> Pi\<^sub>E UNIV X"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1368
    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
  1369
  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
  1370
  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
  1371
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1372
    fix i
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1373
    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
  1374
    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
  1375
      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
  1376
    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
  1377
    define f where "f = min e (1/2)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1378
    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
  1379
    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
  1380
    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
  1381
  } note * = this
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1382
  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
  1383
    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
  1384
  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
  1385
    by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1386
  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
  1387
  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
  1388
    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
  1389
  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
  1390
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1391
    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
  1392
    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
  1393
    proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1394
      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
  1395
        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
  1396
      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
  1397
      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
  1398
        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
  1399
      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
  1400
        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
  1401
      also have "... \<le> (1/2)^(to_nat i) * e i"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1402
        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
  1403
      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
  1404
        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
  1405
      then have "dist (x i) (y i) < e i"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1406
        using \<open>e i < 1\<close> by auto
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1407
      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
  1408
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1409
    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
  1410
    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
  1411
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1412
  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
  1413
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1414
  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
  1415
    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
  1416
  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
  1417
  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
  1418
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1419
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1420
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
  1421
  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
  1422
  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
  1423
  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
  1424
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1425
  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
  1426
    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
  1427
  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
  1428
  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
  1429
  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
  1430
  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
  1431
  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
  1432
    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
  1433
  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
  1434
    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
  1435
  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
  1436
    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
  1437
  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
  1438
    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
  1439
  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
  1440
  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
  1441
    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
  1442
    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
  1443
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1444
  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
  1445
    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
  1446
  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
  1447
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1448
    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
  1449
      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
  1450
      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
  1451
    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
  1452
      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
  1453
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1454
    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
  1455
      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
  1456
    also have "... \<le> 2 * f + e / 8"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1457
      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
  1458
    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
  1459
      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
  1460
    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
  1461
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1462
    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
  1463
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1464
  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
  1465
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1466
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1467
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
  1468
  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
  1469
  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
  1470
proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1471
  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
  1472
  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
  1473
  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
  1474
    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
  1475
next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1476
  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
  1477
  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
  1478
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1479
    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
  1480
    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
  1481
    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
  1482
      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
  1483
    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
  1484
      unfolding K_def using e(2) V(1) V(3) by auto
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1485
    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
  1486
  }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1487
  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
  1488
    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
  1489
  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
  1490
    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
  1491
  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
  1492
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1493
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1494
instance proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1495
  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
  1496
  proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1497
    assume "x = y"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1498
    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
  1499
  next
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1500
    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
  1501
    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
  1502
      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
  1503
    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
  1504
      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
  1505
    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
  1506
      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
  1507
                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
  1508
    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
  1509
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1510
    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
  1511
      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
  1512
    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
  1513
      by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1514
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1515
next
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1516
  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
  1517
        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
  1518
        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
  1519
  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
  1520
    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
  1521
  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
  1522
  {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1523
    fix n
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1524
    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
  1525
            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
  1526
      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
  1527
    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
  1528
      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
  1529
    moreover
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1530
    {
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1531
      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
  1532
      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
  1533
        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
  1534
    }
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1535
    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
  1536
            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
  1537
      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
  1538
  } 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
  1539
  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
  1540
    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
  1541
  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
  1542
                        + (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
  1543
    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
  1544
    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
  1545
      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
  1546
    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
  1547
  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
  1548
                  + (\<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
  1549
    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
  1550
  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
  1551
    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
  1552
  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
  1553
    by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1554
next
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1555
  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
  1556
        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
  1557
        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
  1558
        we copy some mumbo jumbo from Lemma \<open>eventually_uniformity_metric\<close> in
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1559
        \<^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
  1560
  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
  1561
  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
  1562
  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
  1563
    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
  1564
  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
  1565
    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
  1566
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
  1567
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1568
end
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1569
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1570
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
  1571
to first countability, second countability and metrizability, as we have seen above,
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1572
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
  1573
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1574
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
  1575
proof
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1576
  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
  1577
  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
  1578
  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
  1579
    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
  1580
    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
  1581
      using from_nat_to_nat[of i] by metis
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1582
    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
  1583
    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
  1584
      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
  1585
    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
  1586
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1587
    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
  1588
    proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1589
      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
  1590
      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
  1591
              = 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
  1592
        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
  1593
      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
  1594
        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
  1595
        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
  1596
      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
  1597
        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
  1598
      also have "... < (1/2)^k * min e 1"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1599
        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
  1600
      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
  1601
        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
  1602
      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
  1603
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1604
    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
  1605
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1606
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1607
  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
  1608
    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
  1609
  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
  1610
    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
  1611
  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
  1612
  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
  1613
  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
  1614
    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
  1615
    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
  1616
      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
  1617
    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
  1618
      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
  1619
    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
  1620
      by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1621
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1622
    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
  1623
      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
  1624
    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
  1625
    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
  1626
    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
  1627
    proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1628
      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
  1629
      proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1630
        have "K (from_nat n) \<le> L"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1631
          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
  1632
        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
  1633
        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
  1634
      qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1635
      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
  1636
        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
  1637
      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
  1638
        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
  1639
      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
  1640
        apply (rule add_mono)
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1641
        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
  1642
      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
  1643
      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
  1644
    qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1645
    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
  1646
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1647
  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
  1648
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1649
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1650
instance "fun" :: (countable, polish_space) polish_space
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1651
by standard
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1652
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1653
69030
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
  1654
1eb517214deb more on product (function) topologies
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
  1655
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
  1656
subsection \<open>Measurability\<close> (*FIX ME move? *)
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1657
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1658
text \<open>There are two natural sigma-algebras on a product space: the borel sigma algebra,
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1659
generated by open sets in the product, and the product sigma algebra, countably generated by
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1660
products of measurable sets along finitely many coordinates. The second one is defined and studied
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69517
diff changeset
  1661
in \<^file>\<open>Finite_Product_Measure.thy\<close>.
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1662
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1663
These sigma-algebra share a lot of natural properties (measurability of coordinates, for instance),
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1664
but there is a fundamental difference: open sets are generated by arbitrary unions, not only
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1665
countable ones, so typically many open sets will not be measurable with respect to the product sigma
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1666
algebra (while all sets in the product sigma algebra are borel). The two sigma algebras coincide
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1667
only when everything is countable (i.e., the product is countable, and the borel sigma algebra in
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1668
the factor is countably generated).
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1669
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1670
In this paragraph, we develop basic measurability properties for the borel sigma algebra, and
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1671
compare it with the product sigma algebra as explained above.
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1672
\<close>
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1673
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1674
lemma measurable_product_coordinates [measurable (raw)]:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1675
  "(\<lambda>x. x i) \<in> measurable borel borel"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1676
by (rule borel_measurable_continuous_on1[OF continuous_on_product_coordinates])
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1677
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1678
lemma measurable_product_then_coordinatewise:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1679
  fixes f::"'a \<Rightarrow> 'b \<Rightarrow> ('c::topological_space)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1680
  assumes [measurable]: "f \<in> borel_measurable M"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1681
  shows "(\<lambda>x. f x i) \<in> borel_measurable M"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1682
proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1683
  have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1684
    unfolding comp_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
  1685
  then 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
  1686
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1687
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1688
text \<open>To compare the Borel sigma algebra with the product sigma algebra, we give a presentation
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1689
of the product sigma algebra that is more similar to the one we used above for the product
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1690
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
  1691
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1692
lemma sets_PiM_finite:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1693
  "sets (Pi\<^sub>M I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1694
        {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1695
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1696
  have "{(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}} \<subseteq> sets (Pi\<^sub>M I M)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1697
  proof (auto)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1698
    fix X assume H: "\<forall>i. X i \<in> sets (M i)" "finite {i. X i \<noteq> space (M i)}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1699
    then have *: "X i \<in> sets (M i)" for i by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1700
    define J where "J = {i \<in> I. X i \<noteq> space (M i)}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1701
    have "finite J" "J \<subseteq> I" unfolding J_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
  1702
    define Y where "Y = (\<Pi>\<^sub>E j\<in>J. X j)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1703
    have "prod_emb I M J Y \<in> sets (Pi\<^sub>M I M)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1704
      unfolding Y_def apply (rule sets_PiM_I) using \<open>finite 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
  1705
    moreover have "prod_emb I M J Y = (\<Pi>\<^sub>E i\<in>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
  1706
      unfolding prod_emb_def Y_def J_def using H sets.sets_into_space[OF *]
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1707
      by (auto simp add: PiE_iff, blast)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1708
    ultimately show "Pi\<^sub>E I X \<in> sets (Pi\<^sub>M I M)" by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1709
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1710
  then show "sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1711
              \<subseteq> sets (Pi\<^sub>M I M)"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1712
    by (metis (mono_tags, lifting) sets.sigma_sets_subset' sets.top space_PiM)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1713
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1714
  have *: "\<exists>X. {f. (\<forall>i\<in>I. f i \<in> space (M i)) \<and> f \<in> extensional I \<and> f i \<in> A} = Pi\<^sub>E I X \<and>
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1715
                (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1716
    if "i \<in> I" "A \<in> sets (M i)" for i A
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1717
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1718
    define X where "X = (\<lambda>j. if j = i then A else space (M j))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1719
    have "{f. (\<forall>i\<in>I. f i \<in> space (M i)) \<and> f \<in> extensional I \<and> f i \<in> A} = Pi\<^sub>E I X"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1720
      unfolding X_def using sets.sets_into_space[OF \<open>A \<in> sets (M i)\<close>] \<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
  1721
      by (auto simp add: PiE_iff extensional_def, metis subsetCE, metis)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1722
    moreover have "X j \<in> sets (M j)" for j
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1723
      unfolding X_def using \<open>A \<in> sets (M 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
  1724
    moreover have "finite {j. X j \<noteq> space (M j)}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1725
      unfolding X_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
  1726
    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
  1727
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1728
  show "sets (Pi\<^sub>M I M) \<subseteq> sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1729
    unfolding sets_PiM_single
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1730
    apply (rule sigma_sets_mono')
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1731
    apply (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
  1732
    done
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1733
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1734
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1735
lemma sets_PiM_subset_borel:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1736
  "sets (Pi\<^sub>M UNIV (\<lambda>_. borel)) \<subseteq> sets borel"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1737
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1738
  have *: "Pi\<^sub>E UNIV X \<in> sets borel" if [measurable]: "\<And>i. X i \<in> sets borel" "finite {i. X i \<noteq> UNIV}" for X::"'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
  1739
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1740
    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
  1741
    have "finite I" unfolding I_def using that by simp
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1742
    have "Pi\<^sub>E UNIV X = (\<Inter>i\<in>I. (\<lambda>x. x i)-`(X i) \<inter> space borel) \<inter> space borel"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1743
      unfolding 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
  1744
    also have "... \<in> sets borel"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1745
      using that \<open>finite I\<close> by measurable
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1746
    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
  1747
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1748
  then have "{(\<Pi>\<^sub>E i\<in>UNIV. X i) |X::('a \<Rightarrow> 'b set). (\<forall>i. X i \<in> sets borel) \<and> finite {i. X i \<noteq> space borel}} \<subseteq> sets borel"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1749
    by auto
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1750
  then show ?thesis unfolding sets_PiM_finite space_borel
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1751
    by (simp add: * sets.sigma_sets_subset')
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1752
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1753
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1754
proposition sets_PiM_equal_borel:
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1755
  "sets (Pi\<^sub>M UNIV (\<lambda>i::('a::countable). borel::('b::second_countable_topology measure))) = sets borel"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1756
proof
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1757
  obtain K::"('a \<Rightarrow> 'b) set set" where K: "topological_basis K" "countable K"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
  1758
            "\<And>k. k \<in> K \<Longrightarrow> \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV}"
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1759
    using product_topology_countable_basis by fast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1760
  have *: "k \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" 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
  1761
  proof -
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64845
diff changeset
  1762
    obtain X where H: "k = Pi\<^sub>E UNIV X" "\<And>i. open (X i)" "finite {i. X i \<noteq> UNIV}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1763
      using K(3)[OF \<open>k \<in> K\<close>] 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
  1764
    show ?thesis unfolding H(1) sets_PiM_finite space_borel using borel_open[OF H(2)] 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
  1765
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1766
  have **: "U \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" if "open U" for 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
  1767
  proof -
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1768
    obtain B where "B \<subseteq> K" "U = (\<Union>B)"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1769
      using \<open>open U\<close> \<open>topological_basis K\<close> by (metis topological_basis_def)
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1770
    have "countable B" using \<open>B \<subseteq> K\<close> \<open>countable K\<close> countable_subset by blast
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1771
    moreover have "k \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" if "k \<in> B" for k
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1772
      using \<open>B \<subseteq> K\<close> * that by auto
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1773
    ultimately show ?thesis unfolding \<open>U = (\<Union>B)\<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
  1774
  qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1775
  have "sigma_sets UNIV (Collect open) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>i::'a. (borel::('b measure))))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1776
    apply (rule sets.sigma_sets_subset') 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
  1777
  then show "sets (borel::('a \<Rightarrow> 'b) measure) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1778
    unfolding borel_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
  1779
qed (simp add: sets_PiM_subset_borel)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1780
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1781
lemma measurable_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
  1782
  fixes f::"'a \<Rightarrow> ('b::countable) \<Rightarrow> ('c::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
  1783
  assumes [measurable]: "\<And>i. (\<lambda>x. f x i) \<in> borel_measurable M"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1784
  shows "f \<in> borel_measurable M"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1785
proof -
64289
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1786
  have "f \<in> measurable M (Pi\<^sub>M UNIV (\<lambda>_. borel))"
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1787
    by (rule measurable_PiM_single', auto simp add: assms)
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1788
  then show ?thesis using sets_PiM_equal_borel measurable_cong_sets by blast
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1789
qed
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1790
42f28160bad9 HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
hoelzl
parents:
diff changeset
  1791
end