src/HOL/Analysis/Urysohn.thy
author wenzelm
Thu, 25 Jul 2024 10:30:22 +0200
changeset 80616 94703573e0af
parent 78750 f229090cb8a3
permissions -rw-r--r--
tuned names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
     1
(*  Title:      HOL/Analysis/Urysohn.thy
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
    Authors:    LC Paulson, based on material from HOL Light
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
*)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
     5
section \<open>The Urysohn lemma, its consequences and other advanced material about metric spaces\<close>
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Urysohn
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
imports Abstract_Topological_Spaces Abstract_Metric_Spaces Infinite_Sum Arcwise_Connected
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
subsection \<open>Urysohn lemma and Tietze's theorem\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
    13
proposition Urysohn_lemma:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
  fixes a b :: real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
  assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T" "a \<le> b" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  obtains f where "continuous_map X (top_of_set {a..b}) f" "f ` S \<subseteq> {a}" "f ` T \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
  obtain U where "openin X U" "S \<subseteq> U" "X closure_of U \<subseteq> topspace X - T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
    using assms unfolding normal_space_alt disjnt_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
    by (metis Diff_mono Un_Diff_Int closedin_def subset_eq sup_bot_right)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  have "\<exists>G :: real \<Rightarrow> 'a set. G 0 = U \<and> G 1 = topspace X - T \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
               (\<forall>x \<in> dyadics \<inter> {0..1}. \<forall>y \<in> dyadics \<inter> {0..1}. x < y \<longrightarrow> openin X (G x) \<and> openin X (G y) \<and> X closure_of (G x) \<subseteq> G y)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  proof (rule recursion_on_dyadic_fractions)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
    show "openin X U \<and> openin X (topspace X - T) \<and> X closure_of U \<subseteq> topspace X - T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
      using \<open>X closure_of U \<subseteq> topspace X - T\<close> \<open>openin X U\<close> \<open>closedin X T\<close> by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
    show "\<exists>z. (openin X x \<and> openin X z \<and> X closure_of x \<subseteq> z) \<and> openin X z \<and> openin X y \<and> X closure_of z \<subseteq> y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
      if "openin X x \<and> openin X y \<and> X closure_of x \<subseteq> y" for x y
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
      by (meson that closedin_closure_of normal_space_alt \<open>normal_space X\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
    show "openin X x \<and> openin X z \<and> X closure_of x \<subseteq> z"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
      if "openin X x \<and> openin X y \<and> X closure_of x \<subseteq> y" and "openin X y \<and> openin X z \<and> X closure_of y \<subseteq> z" for x y z
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
      by (meson that closure_of_subset openin_subset subset_trans)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  then obtain G :: "real \<Rightarrow> 'a set"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
      where G0: "G 0 = U" and G1: "G 1 = topspace X - T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
        and G: "\<And>x y. \<lbrakk>x \<in> dyadics; y \<in> dyadics; 0 \<le> x; x < y; y \<le> 1\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
                      \<Longrightarrow> openin X (G x) \<and> openin X (G y) \<and> X closure_of (G x) \<subseteq> G y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
    by (smt (verit, del_insts) Int_iff atLeastAtMost_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  define f where "f \<equiv> \<lambda>x. Inf(insert 1 {r. r \<in> dyadics \<inter> {0..1} \<and> x \<in> G r})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
  have f_ge: "f x \<ge> 0" if "x \<in> topspace X" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
    unfolding f_def by (force intro: cInf_greatest)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  moreover have f_le1: "f x \<le> 1" if "x \<in> topspace X" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    have "bdd_below {r \<in> dyadics \<inter> {0..1}. x \<in> G r}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
      by (auto simp: bdd_below_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
       by (auto simp: f_def cInf_lower)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  ultimately have fim: "f ` topspace X \<subseteq> {0..1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    by (auto simp: f_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  have 0: "0 \<in> dyadics \<inter> {0..1::real}" and 1: "1 \<in> dyadics \<inter> {0..1::real}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
    by (force simp: dyadics_def)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  then have opeG: "openin X (G r)" if "r \<in> dyadics \<inter> {0..1}" for r
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
    using G G0 \<open>openin X U\<close> less_eq_real_def that by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  have "x \<in> G 0" if "x \<in> S" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
    using G0 \<open>S \<subseteq> U\<close> that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
  with 0 have fimS: "f ` S \<subseteq> {0}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
    unfolding f_def by (force intro!: cInf_eq_minimum)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  have False if "r \<in> dyadics" "0 \<le> r" "r < 1" "x \<in> G r" "x \<in> T" for r x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
    using G [of r 1] 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
    by (smt (verit, best) DiffD2 G1 Int_iff closure_of_subset inf.orderE openin_subset that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  then have "r\<ge>1" if "r \<in> dyadics" "0 \<le> r" "r \<le> 1" "x \<in> G r" "x \<in> T" for r x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
    using linorder_not_le that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  then have fimT: "f ` T \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
    unfolding f_def by (force intro!: cInf_eq_minimum)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  have fle1: "f z \<le> 1" for z
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
    by (force simp: f_def intro: cInf_lower)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  have fle: "f z \<le> x" if "x \<in> dyadics \<inter> {0..1}" "z \<in> G x" for z x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    using that by (force simp: f_def intro: cInf_lower)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  have *: "b \<le> f z" if "b \<le> 1" "\<And>x. \<lbrakk>x \<in> dyadics \<inter> {0..1}; z \<in> G x\<rbrakk> \<Longrightarrow> b \<le> x" for z b
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    using that by (force simp: f_def intro: cInf_greatest)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
  have **: "r \<le> f x" if r: "r \<in> dyadics \<inter> {0..1}" "x \<notin> G r" for r x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
  proof (rule *)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
    show "r \<le> s" if "s \<in> dyadics \<inter> {0..1}" and "x \<in> G s" for s :: real
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
    74
      using that r G [of s r] by (force simp: dest: closure_of_subset openin_subset)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
  qed (use that in force)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  have "\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y \<in> U. \<bar>f y - f x\<bar> < \<epsilon>)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
    if "x \<in> topspace X" and "0 < \<epsilon>" for x \<epsilon>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
    have A: "\<exists>r. r \<in> dyadics \<inter> {0..1} \<and> r < y \<and> \<bar>r - y\<bar> < d" if "0 < y" "y \<le> 1" "0 < d" for y d::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
      obtain n q r 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
        where "of_nat q / 2^n < y" "y < of_nat r / 2^n" "\<bar>q / 2^n - r / 2^n \<bar> < d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
        by (smt (verit, del_insts) padic_rational_approximation_straddle_pos  \<open>0 < d\<close> \<open>0 < y\<close>) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
        unfolding dyadics_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
        using divide_eq_0_iff that(2) by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
    have B: "\<exists>r. r \<in> dyadics \<inter> {0..1} \<and> y < r \<and> \<bar>r - y\<bar> < d" if "0 \<le> y" "y < 1" "0 < d" for y d::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
      obtain n q r 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
        where "of_nat q / 2^n \<le> y" "y < of_nat r / 2^n" "\<bar>q / 2^n - r / 2^n \<bar> < d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
        using padic_rational_approximation_straddle_pos_le
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
        by (smt (verit, del_insts) \<open>0 < d\<close> \<open>0 \<le> y\<close>) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
        apply (clarsimp simp: dyadics_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
        using divide_eq_0_iff \<open>y < 1\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
        by (smt (verit) divide_nonneg_nonneg divide_self of_nat_0_le_iff of_nat_1 power_0 zero_le_power) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
    proof (cases "f x = 0")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
      with B[of 0] obtain r where r: "r \<in> dyadics \<inter> {0..1}" "0 < r" "\<bar>r\<bar> < \<epsilon>/2"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
        by (smt (verit) \<open>0 < \<epsilon>\<close> half_gt_zero)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
      proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
        show "openin X (G r)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
          using opeG r(1) by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
        show "x \<in> G r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
          using True ** r by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
        show "\<forall>y \<in> G r. \<bar>f y - f x\<bar> < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
          using f_ge \<open>openin X (G r)\<close> fle openin_subset r by (fastforce simp: True)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
      show ?thesis 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
      proof (cases "f x = 1")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
        case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
        with A[of 1] obtain r where r: "r \<in> dyadics \<inter> {0..1}" "r < 1" "\<bar>r-1\<bar> < \<epsilon>/2"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
          by (smt (verit) \<open>0 < \<epsilon>\<close> half_gt_zero)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
        define G' where "G' \<equiv> topspace X - X closure_of G r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
        show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
        proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
          show "openin X G'"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
            unfolding G'_def by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
          obtain r' where "r' \<in> dyadics \<and> 0 \<le> r' \<and> r' \<le> 1 \<and> r < r' \<and> \<bar>r' - r\<bar> < 1 - r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
            using B r by force 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
          moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
          have "1 - r \<in> dyadics" "0 \<le> r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
            using 1 r dyadics_diff by force+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
          ultimately have "x \<notin> X closure_of G r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
            using G True r fle by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
          then show "x \<in> G'"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
            by (simp add: G'_def that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
          show "\<forall>y \<in> G'. \<bar>f y - f x\<bar> < \<epsilon>"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   136
            using ** f_le1 in_closure_of r by (fastforce simp: True G'_def)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
      next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
        case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
        have "0 < f x" "f x < 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
          using fle1 f_ge that(1) \<open>f x \<noteq> 0\<close> \<open>f x \<noteq> 1\<close> by (metis order_le_less) +
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
        obtain r where r: "r \<in> dyadics \<inter> {0..1}" "r < f x" "\<bar>r - f x\<bar> < \<epsilon> / 2"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
          using A \<open>0 < \<epsilon>\<close> \<open>0 < f x\<close> \<open>f x < 1\<close> by (smt (verit, best) half_gt_zero)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
        obtain r' where r': "r' \<in> dyadics \<inter> {0..1}" "f x < r'" "\<bar>r' - f x\<bar> < \<epsilon> / 2"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
          using B \<open>0 < \<epsilon>\<close> \<open>0 < f x\<close> \<open>f x < 1\<close> by (smt (verit, best) half_gt_zero)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
        have "r < 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
          using \<open>f x < 1\<close> r(2) by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
        show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
        proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
          show "openin X (G r' - X closure_of G r)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
            using closedin_closure_of opeG r' by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
          have "x \<in> X closure_of G r \<Longrightarrow> False"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
            using B [of r "f x - r"] r \<open>r < 1\<close> G [of r] fle by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
          then show "x \<in> G r' - X closure_of G r"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
            using ** r' by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
          show "\<forall>y\<in>G r' - X closure_of G r. \<bar>f y - f x\<bar> < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
            using r r' ** G closure_of_subset field_sum_of_halves fle openin_subset subset_eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
            by (smt (verit) DiffE opeG)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  then have contf: "continuous_map X (top_of_set {0..1}) f"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   164
    by (force simp: Met_TC.continuous_map_to_metric dist_real_def continuous_map_in_subtopology fim simp flip: mtopology_is_euclidean)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  define g where "g \<equiv> \<lambda>x. a + (b - a) * f x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
  show thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
    have "continuous_map X euclideanreal g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
      using contf \<open>a \<le> b\<close> unfolding g_def by (auto simp: continuous_intros continuous_map_in_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
    moreover have "g ` (topspace X) \<subseteq> {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
      using mult_left_le [of "f _" "b-a"] contf \<open>a \<le> b\<close>   
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
      by (simp add: g_def add.commute continuous_map_in_subtopology image_subset_iff le_diff_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
    ultimately show "continuous_map X (top_of_set {a..b}) g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
      by (meson continuous_map_in_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
    show "g ` S \<subseteq> {a}" "g ` T \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
      using fimS fimT by (auto simp: g_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
lemma Urysohn_lemma_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  fixes a b :: real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  assumes "normal_space X" "closedin X S" "closedin X T" "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  obtains f where "continuous_map X euclideanreal f" "f ` S \<subseteq> {a}" "f ` T \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
  by (metis Urysohn_lemma assms continuous_map_in_subtopology disjnt_sym linear)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
lemma normal_space_iff_Urysohn_gen_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  assumes "a \<noteq> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
  shows "normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
         (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
                \<longrightarrow> (\<exists>f. continuous_map X euclideanreal f \<and> f ` S \<subseteq> {a} \<and> f ` T \<subseteq> {b}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
 (is "?lhs=?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  show "?lhs \<Longrightarrow> ?rhs" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
    by (metis Urysohn_lemma_alt)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
  assume R: ?rhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
    unfolding normal_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    assume "closedin X S" and "closedin X T" and "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
    with R obtain f where contf: "continuous_map X euclideanreal f" and "f ` S \<subseteq> {a}" "f ` T \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
      by meson
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    show "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
    proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
      show "openin X {x \<in> topspace X. f x \<in> ball a (\<bar>a - b\<bar> / 2)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
        by (force intro!: openin_continuous_map_preimage [OF contf])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
      show "openin X {x \<in> topspace X. f x \<in> ball b (\<bar>a - b\<bar> / 2)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
        by (force intro!: openin_continuous_map_preimage [OF contf])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
      show "S \<subseteq> {x \<in> topspace X. f x \<in> ball a (\<bar>a - b\<bar> / 2)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
        using \<open>closedin X S\<close> closedin_subset \<open>f ` S \<subseteq> {a}\<close> assms by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
      show "T \<subseteq> {x \<in> topspace X. f x \<in> ball b (\<bar>a - b\<bar> / 2)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
        using \<open>closedin X T\<close> closedin_subset \<open>f ` T \<subseteq> {b}\<close> assms by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
      have "\<And>x. \<lbrakk>x \<in> topspace X; dist a (f x) < \<bar>a-b\<bar>/2; dist b (f x) < \<bar>a-b\<bar>/2\<rbrakk> \<Longrightarrow> False"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
        by (smt (verit, best) dist_real_def dist_triangle_half_l)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
      then show "disjnt {x \<in> topspace X. f x \<in> ball a (\<bar>a-b\<bar> / 2)} {x \<in> topspace X. f x \<in> ball b (\<bar>a-b\<bar> / 2)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
        using disjnt_iff by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
qed 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
lemma normal_space_iff_Urysohn_gen:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
  shows
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
   "a < b \<Longrightarrow> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
      normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
        (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
               \<longrightarrow> (\<exists>f. continuous_map X (top_of_set {a..b}) f \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
                        f ` S \<subseteq> {a} \<and> f ` T \<subseteq> {b}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  by (metis linear not_le Urysohn_lemma normal_space_iff_Urysohn_gen_alt continuous_map_in_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
lemma normal_space_iff_Urysohn_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
   "normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
     (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
           \<longrightarrow> (\<exists>f. continuous_map X euclideanreal f \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
                   f ` S \<subseteq> {0} \<and> f ` T \<subseteq> {1}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
  by (rule normal_space_iff_Urysohn_gen_alt) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
lemma normal_space_iff_Urysohn:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
   "normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
     (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
            \<longrightarrow> (\<exists>f::'a\<Rightarrow>real. continuous_map X (top_of_set {0..1}) f \<and> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
                               f ` S \<subseteq> {0} \<and> f ` T \<subseteq> {1}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
  by (rule normal_space_iff_Urysohn_gen) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
lemma normal_space_perfect_map_image:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
   "\<lbrakk>normal_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> normal_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
  unfolding perfect_map_def proper_map_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
  using normal_space_continuous_closed_map_image by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
lemma Hausdorff_normal_space_closed_continuous_map_image:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
   "\<lbrakk>normal_space X; closed_map X Y f; continuous_map X Y f;
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
     f ` topspace X = topspace Y; t1_space Y\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
    \<Longrightarrow> Hausdorff_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
  by (metis normal_space_continuous_closed_map_image normal_t1_imp_Hausdorff_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
lemma normal_Hausdorff_space_closed_continuous_map_image:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
   "\<lbrakk>normal_space X; Hausdorff_space X; closed_map X Y f;
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
     continuous_map X Y f; f ` topspace X = topspace Y\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
    \<Longrightarrow> normal_space Y \<and> Hausdorff_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
  by (meson normal_space_continuous_closed_map_image normal_t1_eq_Hausdorff_space t1_space_closed_map_image)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
lemma Lindelof_cover:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
  assumes "regular_space X" and "Lindelof_space X" and "S \<noteq> {}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
    and clo: "closedin X S" "closedin X T" "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  obtains h :: "nat \<Rightarrow> 'a set" where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    "\<And>n. openin X (h n)" "\<And>n. disjnt T (X closure_of (h n))" and  "S \<subseteq> \<Union>(range h)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
  have "\<exists>U. openin X U \<and> x \<in> U \<and> disjnt T (X closure_of U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
    if "x \<in> S" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    using \<open>regular_space X\<close> unfolding regular_space 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
    by (metis (full_types) Diff_iff \<open>disjnt S T\<close> clo closure_of_eq disjnt_iff in_closure_of that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
  then obtain h where oh: "\<And>x. x \<in> S \<Longrightarrow> openin X (h x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
    and xh: "\<And>x. x \<in> S \<Longrightarrow> x \<in> h x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
    and dh: "\<And>x. x \<in> S \<Longrightarrow> disjnt T (X closure_of h x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
  have "Lindelof_space(subtopology X S)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    by (simp add: Lindelof_space_closedin_subtopology \<open>Lindelof_space X\<close> \<open>closedin X S\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
  then obtain \<U> where \<U>: "countable \<U> \<and> \<U> \<subseteq> h ` S \<and> S \<subseteq> \<Union>\<U>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
    unfolding Lindelof_space_subtopology_subset [OF closedin_subset [OF \<open>closedin X S\<close>]]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
    by (smt (verit, del_insts) oh xh UN_I image_iff subsetI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
  with \<open>S \<noteq> {}\<close> have "\<U> \<noteq> {}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
    by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
  proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
    show "openin X (from_nat_into \<U> n)" for n
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
      by (metis \<U> from_nat_into image_iff \<open>\<U> \<noteq> {}\<close> oh subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
    show "disjnt T (X closure_of (from_nat_into \<U>) n)" for n
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
      using dh from_nat_into [OF \<open>\<U> \<noteq> {}\<close>]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
      by (metis \<U> f_inv_into_f inv_into_into subset_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
    show "S \<subseteq> \<Union>(range (from_nat_into \<U>))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
      by (simp add: \<U> \<open>\<U> \<noteq> {}\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
lemma regular_Lindelof_imp_normal_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
  assumes "regular_space X" and "Lindelof_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
  shows "normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
  unfolding normal_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
  fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
  assume clo: "closedin X S" "closedin X T" and "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
  show "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
  proof (cases "S={} \<or> T={}")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
    case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
    with clo show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
      by (meson closedin_def disjnt_empty1 disjnt_empty2 openin_empty openin_topspace subset_empty)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
    obtain h :: "nat \<Rightarrow> 'a set" where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
      opeh: "\<And>n. openin X (h n)" and dish: "\<And>n. disjnt T (X closure_of (h n))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
      and Sh: "S \<subseteq> \<Union>(range h)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
      by (metis Lindelof_cover False \<open>disjnt S T\<close> assms clo)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
    obtain k :: "nat \<Rightarrow> 'a set" where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
      opek: "\<And>n. openin X (k n)" and disk: "\<And>n. disjnt S (X closure_of (k n))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
      and Tk: "T \<subseteq> \<Union>(range k)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      by (metis Lindelof_cover False \<open>disjnt S T\<close> assms clo disjnt_sym)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
    define U where "U \<equiv> \<Union>i. h i - (\<Union>j<i. X closure_of k j)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
    define V where "V \<equiv> \<Union>i. k i - (\<Union>j\<le>i. X closure_of h j)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
    proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
      show "openin X U" "openin X V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
        unfolding U_def V_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
        by (force intro!: opek opeh closedin_Union closedin_closure_of)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
      show "S \<subseteq> U" "T \<subseteq> V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
        using Sh Tk dish disk by (fastforce simp: U_def V_def disjnt_iff)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
      have "\<And>x i j. \<lbrakk>x \<in> k i; x \<in> h j; \<forall>j\<le>i. x \<notin> X closure_of h j\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
                 \<Longrightarrow> \<exists>i<j. x \<in> X closure_of k i"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
        by (metis in_closure_of linorder_not_less opek openin_subset subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
      then show "disjnt U V"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   331
        by (force simp: U_def V_def disjnt_iff)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
   336
theorem Tietze_extension_closed_real_interval:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
  assumes "normal_space X" and "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
    and contf: "continuous_map (subtopology X S) euclideanreal f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
    and fim: "f ` S \<subseteq> {a..b}" and "a \<le> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
  obtains g 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
  where "continuous_map X euclideanreal g" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
        "\<And>x. x \<in> S \<Longrightarrow> g x = f x" "g ` topspace X \<subseteq> {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
  define c where "c \<equiv> max \<bar>a\<bar> \<bar>b\<bar> + 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
  have "0 < c" and c: "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<le> c"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
    using fim by (auto simp: c_def image_subset_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
  define good where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
    "good \<equiv> \<lambda>g n. continuous_map X euclideanreal g \<and> (\<forall>x \<in> S. \<bar>f x - g x\<bar> \<le> c * (2/3)^n)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
  have step: "\<exists>g. good g (Suc n) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
              (\<forall>x \<in> topspace X. \<bar>g x - h x\<bar> \<le> c * (2/3)^n / 3)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
    if h: "good h n" for n h
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
    have pos: "0 < c * (2/3) ^ n"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
      by (simp add: \<open>0 < c\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
    have S_eq: "S = topspace(subtopology X S)" and "S \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
      using \<open>closedin X S\<close> closedin_subset by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
    define d where "d \<equiv> c/3 * (2/3) ^ n"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
    define SA where "SA \<equiv> {x \<in> S. f x - h x \<in> {..-d}}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
    define SB where "SB \<equiv> {x \<in> S. f x - h x \<in> {d..}}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
    have contfh: "continuous_map (subtopology X S) euclideanreal (\<lambda>x. f x - h x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
      using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
      by (simp add: contf good_def continuous_map_diff continuous_map_from_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
    then have "closedin (subtopology X S) SA"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
      unfolding SA_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
      by (smt (verit, del_insts) closed_closedin continuous_map_closedin Collect_cong S_eq closed_real_atMost)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
    then have "closedin X SA"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
      using \<open>closedin X S\<close> closedin_trans_full by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
    moreover have  "closedin (subtopology X S) SB"      
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
      unfolding SB_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
      using closedin_continuous_map_preimage_gen [OF contfh]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
      by (metis (full_types) S_eq closed_atLeast closed_closedin closedin_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
    then have "closedin X SB"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
      using \<open>closedin X S\<close> closedin_trans_full by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
    moreover have "disjnt SA SB"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
      using pos by (auto simp: d_def disjnt_def SA_def SB_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
    moreover have "-d \<le> d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
      using pos by (auto simp: d_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
    ultimately
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
    obtain g where contg: "continuous_map X (top_of_set {- d..d}) g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
      and ga: "g ` SA \<subseteq> {- d}" and gb: "g ` SB \<subseteq> {d}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
      using Urysohn_lemma \<open>normal_space X\<close> by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
    then have g_le_d: "\<And>x. x \<in> topspace X \<Longrightarrow> \<bar>g x\<bar> \<le> d"
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78283
diff changeset
   383
      by (fastforce simp: abs_le_iff continuous_map_def minus_le_iff)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
    have g_eq_d: "\<And>x. \<lbrakk>x \<in> S; f x - h x \<le> -d\<rbrakk> \<Longrightarrow> g x = -d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
      using ga by (auto simp: SA_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
    have g_eq_negd: "\<And>x. \<lbrakk>x \<in> S; f x - h x \<ge> d\<rbrakk> \<Longrightarrow> g x = d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
      using gb by (auto simp: SB_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
      unfolding good_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
    proof (intro conjI strip exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
      show "continuous_map X euclideanreal (\<lambda>x. h x + g x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
        using contg continuous_map_add continuous_map_in_subtopology that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
        unfolding good_def by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
      show "\<bar>f x - (h x + g x)\<bar> \<le> c * (2 / 3) ^ Suc n" if "x \<in> S" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
      proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
        have x: "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
          using \<open>S \<subseteq> topspace X\<close> that by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
        have "\<bar>f x - h x\<bar> \<le> c * (2/3) ^ n"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
          using good_def h that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
        with g_eq_d [OF that] g_eq_negd [OF that] g_le_d [OF x] 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
        have "\<bar>f x - (h x + g x)\<bar> \<le> d + d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
          unfolding d_def by linarith
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
        then show ?thesis 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
          by (simp add: d_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
      show "\<bar>h x + g x - h x\<bar> \<le> c * (2 / 3) ^ n / 3" if "x \<in> topspace X" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
        using that d_def g_le_d by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
  then obtain nxtg where nxtg: "\<And>h n. good h n \<Longrightarrow> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
          good (nxtg h n) (Suc n) \<and> (\<forall>x \<in> topspace X. \<bar>nxtg h n x - h x\<bar> \<le> c * (2/3)^n / 3)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
    by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
  define g where "g \<equiv> rec_nat (\<lambda>x. 0) (\<lambda>n r. nxtg r n)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
  have [simp]: "g 0 x = 0" for x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
    by (auto simp: g_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
  have g_Suc: "g(Suc n) = nxtg (g n) n" for n
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
    by (auto simp: g_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
  have good: "good (g n) n" for n
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
  proof (induction n)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
    case 0
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
    with c show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
      by (auto simp: good_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
  qed (simp add: g_Suc nxtg)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
  have *: "\<And>n x. x \<in> topspace X \<Longrightarrow> \<bar>g(Suc n) x - g n x\<bar> \<le> c * (2/3) ^ n / 3"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
    using nxtg g_Suc good by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
  obtain h where conth:  "continuous_map X euclideanreal h"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
    and h: "\<And>\<epsilon>. 0 < \<epsilon> \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>topspace X. dist (g n x) (h x) < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
  proof (rule Met_TC.continuous_map_uniformly_Cauchy_limit)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
    show "\<forall>\<^sub>F n in sequentially. continuous_map X (Met_TC.mtopology) (g n)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
      using good good_def by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
    show "\<exists>N. \<forall>m n x. N \<le> m \<longrightarrow> N \<le> n \<longrightarrow> x \<in> topspace X \<longrightarrow> dist (g m x) (g n x) < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
      if "\<epsilon> > 0" for \<epsilon> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
      have "\<forall>\<^sub>F n in sequentially. \<bar>(2/3) ^ n\<bar> < \<epsilon>/c"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
      proof (rule Archimedean_eventually_pow_inverse)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
        show "0 < \<epsilon> / c"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
          by (simp add: \<open>0 < c\<close> that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
      qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
      then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> \<bar>(2/3) ^ n\<bar> < \<epsilon>/c"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
        by (meson eventually_sequentially order_le_less_trans)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
      have "\<bar>g m x - g n x\<bar> < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
        if "N \<le> m" "N \<le> n" and x: "x \<in> topspace X" "m \<le> n" for m n x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
      proof (cases "m < n")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
        case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
        have 23: "(\<Sum>k = m..<n. (2/3)^k) = 3 * ((2/3) ^ m - (2/3::real) ^ n)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
          using \<open>m \<le> n\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
          by (induction n) (auto simp: le_Suc_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
        have "\<bar>g m x - g n x\<bar> \<le> \<bar>\<Sum>k = m..<n. g (Suc k) x - g k x\<bar>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
          by (subst sum_Suc_diff' [OF \<open>m \<le> n\<close>]) linarith
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
        also have "\<dots> \<le> (\<Sum>k = m..<n. \<bar>g (Suc k) x - g k x\<bar>)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
          by (rule sum_abs)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
        also have "\<dots> \<le> (\<Sum>k = m..<n. c * (2/3)^k / 3)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
          by (meson "*" sum_mono x(1))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
        also have "\<dots> = (c/3) * (\<Sum>k = m..<n. (2/3)^k)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
          by (simp add: sum_distrib_left)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
        also have "\<dots> = (c/3) * 3 * ((2/3) ^ m - (2/3) ^ n)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
          by (simp add: sum_distrib_left 23)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
        also have "... < (c/3) * 3 * ((2/3) ^ m)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
          using \<open>0 < c\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
        also have "\<dots> < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
          using N [OF \<open>N \<le> m\<close>] \<open>0 < c\<close> by (simp add: field_simps)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
        finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
      qed (use \<open>0 < \<epsilon>\<close> \<open>m \<le> n\<close> in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
        by (metis dist_commute_lessI dist_real_def nle_le)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
  qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
  define \<phi> where "\<phi> \<equiv> \<lambda>x. max a (min (h x) b)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
  show thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
  proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
    show "continuous_map X euclidean \<phi>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
      unfolding \<phi>_def using conth by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
    show "\<phi> x = f x" if "x \<in> S" for x 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
      have x: "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
        using \<open>closedin X S\<close> closedin_subset that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
      have "h x = f x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
      proof (rule Met_TC.limitin_metric_unique)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
        show "limitin Met_TC.mtopology (\<lambda>n. g n x) (h x) sequentially"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
          using h x by (force simp: tendsto_iff eventually_sequentially)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
        show "limitin Met_TC.mtopology (\<lambda>n. g n x) (f x) sequentially"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
        proof (clarsimp simp: tendsto_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
          fix \<epsilon>::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
          assume "\<epsilon> > 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
          then have "\<forall>\<^sub>F n in sequentially. \<bar>(2/3) ^ n\<bar> < \<epsilon>/c"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
            by (intro Archimedean_eventually_pow_inverse) (auto simp: \<open>c > 0\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
          then show "\<forall>\<^sub>F n in sequentially. dist (g n x) (f x) < \<epsilon>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
            apply eventually_elim
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
            by (smt (verit) good x good_def \<open>c > 0\<close> dist_real_def mult.commute pos_less_divide_eq that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
      qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
        using that fim by (auto simp: \<phi>_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
    then show "\<phi> ` topspace X \<subseteq> {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
      using fim \<open>a \<le> b\<close> by (auto simp: \<phi>_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
   501
theorem Tietze_extension_realinterval:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
  assumes XS: "normal_space X" "closedin X S" and T: "is_interval T" "T \<noteq> {}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
    and contf: "continuous_map (subtopology X S) euclideanreal f" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
    and "f ` S \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
  obtains g where "continuous_map X euclideanreal g"  "g ` topspace X \<subseteq> T"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
  define \<Phi> where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
        "\<Phi> \<equiv> \<lambda>T::real set. \<forall>f. continuous_map (subtopology X S) euclidean f \<longrightarrow> f`S \<subseteq> T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
               \<longrightarrow> (\<exists>g. continuous_map X euclidean g \<and> g ` topspace X \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
  have "\<Phi> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
    if *: "\<And>T. \<lbrakk>bounded T; is_interval T; T \<noteq> {}\<rbrakk> \<Longrightarrow> \<Phi> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
      and "is_interval T" "T \<noteq> {}" for T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
    unfolding \<Phi>_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
    fix f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
    assume contf: "continuous_map (subtopology X S) euclideanreal f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
      and "f ` S \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
    have \<Phi>T: "\<Phi> ((\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
    proof (rule *)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
      show "bounded ((\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
        using shrink_range [of T] by (force intro: boundedI [where B=1])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
      show "is_interval ((\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
        using connected_shrink that(2) is_interval_connected_1 by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
      show "(\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T \<noteq> {}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
        using \<open>T \<noteq> {}\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
    moreover have "continuous_map (subtopology X S) euclidean ((\<lambda>x. x / (1 + \<bar>x\<bar>)) \<circ> f)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
      by (metis contf continuous_map_compose continuous_map_into_fulltopology continuous_map_real_shrink)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
    moreover have "((\<lambda>x. x / (1 + \<bar>x\<bar>)) \<circ> f) ` S \<subseteq> (\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
      using \<open>f ` S \<subseteq> T\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
    ultimately obtain g 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
       where contg: "continuous_map X euclidean g" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
         and gim: "g ` topspace X \<subseteq> (\<lambda>x. x / (1 + \<bar>x\<bar>)) ` T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
         and geq: "\<And>x. x \<in> S \<Longrightarrow> g x = ((\<lambda>x. x / (1 + \<bar>x\<bar>)) \<circ> f) x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
      using \<Phi>T unfolding \<Phi>_def by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
    show "\<exists>g. continuous_map X euclideanreal g \<and> g ` topspace X \<subseteq> T \<and> (\<forall>x\<in>S. g x = f x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
      have "continuous_map X (top_of_set {-1<..<1}) g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
        using contg continuous_map_in_subtopology gim shrink_range by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
      then show "continuous_map X euclideanreal ((\<lambda>x. x / (1 - \<bar>x\<bar>)) \<circ> g)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
        by (rule continuous_map_compose) (auto simp: continuous_on_real_grow)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
      show "((\<lambda>x. x / (1 - \<bar>x\<bar>)) \<circ> g) ` topspace X \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
        using gim real_grow_shrink by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
      show "\<forall>x\<in>S. ((\<lambda>x. x / (1 - \<bar>x\<bar>)) \<circ> g) x = f x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
        using geq real_grow_shrink by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
  moreover have "\<Phi> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
    if "bounded T" "is_interval T" "T \<noteq> {}" for T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
    unfolding \<Phi>_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
    fix f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
    assume contf: "continuous_map (subtopology X S) euclideanreal f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
      and "f ` S \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
    obtain a b where ab: "closure T = {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
      by (meson \<open>bounded T\<close> \<open>is_interval T\<close> compact_closure connected_compact_interval_1 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
            connected_imp_connected_closure is_interval_connected)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
    with \<open>T \<noteq> {}\<close> have "a \<le> b" by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
    have "f ` S \<subseteq> {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
      using \<open>f ` S \<subseteq> T\<close> ab closure_subset by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
    then obtain g where contg: "continuous_map X euclideanreal g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
      and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x" and gim: "g ` topspace X \<subseteq> {a..b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
      using Tietze_extension_closed_real_interval [OF XS contf _ \<open>a \<le> b\<close>] by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
    define W where "W \<equiv> {x \<in> topspace X. g x \<in> closure T - T}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
    have "{a..b} - {a, b} \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
      using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
      by (metis ab atLeastAtMost_diff_ends convex_interior_closure interior_atLeastAtMost_real 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
          interior_subset is_interval_convex)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
    with finite_imp_compact have "compact (closure T - T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
      by (metis Diff_eq_empty_iff Diff_insert2 ab finite.emptyI finite_Diff_insert)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
    then have "closedin X W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
      unfolding W_def using closedin_continuous_map_preimage [OF contg] compact_imp_closed by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
    moreover have "disjnt W S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
      unfolding W_def disjnt_iff using \<open>f ` S \<subseteq> T\<close> gf by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
    ultimately obtain h :: "'a \<Rightarrow> real" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
      where conth: "continuous_map X (top_of_set {0..1}) h" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
            and him: "h ` W \<subseteq> {0}" "h ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
      by (metis XS normal_space_iff_Urysohn) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
    then have him01: "h ` topspace X \<subseteq> {0..1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
      by (meson continuous_map_in_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
    obtain z where "z \<in> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
      using \<open>T \<noteq> {}\<close> by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
    define g' where "g' \<equiv> \<lambda>x. z + h x * (g x - z)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
    show "\<exists>g. continuous_map X euclidean g \<and> g ` topspace X \<subseteq> T \<and> (\<forall>x\<in>S. g x = f x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
    proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
      show "continuous_map X euclideanreal g'"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
        unfolding g'_def using contg conth continuous_map_in_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
        by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
      show "g' ` topspace X \<subseteq> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
        unfolding g'_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
      proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
        fix x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
        assume "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
        show "z + h x * (g x - z) \<in> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
        proof (cases "g x \<in> T")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
          case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
          define w where "w \<equiv> z + h x * (g x - z)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
          have "\<bar>h x\<bar> * \<bar>g x - z\<bar> \<le> \<bar>g x - z\<bar>" "\<bar>1 - h x\<bar> * \<bar>g x - z\<bar> \<le> \<bar>g x - z\<bar>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
            using him01 \<open>x \<in> topspace X\<close> by (force simp: intro: mult_left_le_one_le)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
          then consider "z \<le> w \<and> w \<le> g x" | "g x \<le> w \<and> w \<le> z"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
            unfolding w_def by (smt (verit) left_diff_distrib mult_cancel_right2 mult_minus_right zero_less_mult_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
            using \<open>is_interval T\<close> unfolding w_def is_interval_1 by (metis True \<open>z \<in> T\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
        next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
          case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
          then have "g x \<in> closure T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
            using \<open>x \<in> topspace X\<close> ab gim by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
          then have "h x = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
            using him False \<open>x \<in> topspace X\<close> by (auto simp: W_def image_subset_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
            by (simp add: \<open>z \<in> T\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
      show "\<forall>x\<in>S. g' x = f x"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
        using gf him by (auto simp: W_def g'_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
    qed 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
  ultimately show thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
    using assms that unfolding \<Phi>_def by best
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
lemma normal_space_iff_Tietze:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
   "normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
    (\<forall>f S. closedin X S \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
           continuous_map (subtopology X S) euclidean f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
           \<longrightarrow> (\<exists>g:: 'a \<Rightarrow> real. continuous_map X euclidean g \<and> (\<forall>x \<in> S. g x = f x)))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
  assume ?lhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
  then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
    by (metis Tietze_extension_realinterval empty_not_UNIV is_interval_univ subset_UNIV)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
  assume R: ?rhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
    unfolding normal_space_iff_Urysohn_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
    assume "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
      and "closedin X T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
      and "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
    then have cloST: "closedin X (S \<union> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
      by (simp add: closedin_Un)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
    moreover 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
    have "continuous_map (subtopology X (S \<union> T)) euclideanreal (\<lambda>x. if x \<in> S then 0 else 1)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
      unfolding continuous_map_closedin
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
    proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
      fix C :: "real set"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
      define D where "D \<equiv> {x \<in> topspace X. if x \<in> S then 0 \<in> C else x \<in> T \<and> 1 \<in> C}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
      have "D \<in> {{}, S, T, S \<union> T}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
        unfolding D_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
        using closedin_subset [OF \<open>closedin X S\<close>] closedin_subset [OF \<open>closedin X T\<close>] \<open>disjnt S T\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
        by (auto simp: disjnt_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
      then have "closedin X D"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
        using \<open>closedin X S\<close> \<open>closedin X T\<close> closedin_empty by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
      with closedin_subset_topspace
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
      show "closedin (subtopology X (S \<union> T)) {x \<in> topspace (subtopology X (S \<union> T)). (if x \<in> S then 0::real else 1) \<in> C}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
        apply (simp add: D_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
        by (smt (verit, best) Collect_cong Collect_mono_iff Un_def closedin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
    qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
    ultimately obtain g :: "'a \<Rightarrow> real"  where 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
      contg: "continuous_map X euclidean g" and gf: "\<forall>x \<in> S \<union> T. g x = (if x \<in> S then 0 else 1)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
      using R by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
    then show "\<exists>f. continuous_map X euclideanreal f \<and> f ` S \<subseteq> {0} \<and> f ` T \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
      by (smt (verit) Un_iff \<open>disjnt S T\<close> disjnt_iff image_subset_iff insert_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
   668
subsection \<open>Random metric space stuff\<close>
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
lemma metrizable_imp_k_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
  assumes "metrizable_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
  shows "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
  obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
    using assms unfolding metrizable_space_def by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
  then interpret Metric_space M d 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
    by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
    unfolding k_space Xeq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
  proof clarsimp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
    fix S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
    assume "S \<subseteq> M" and S: "\<forall>K. compactin mtopology K \<longrightarrow> closedin (subtopology mtopology K) (K \<inter> S)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
    have "l \<in> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
      if \<sigma>: "range \<sigma> \<subseteq> S" and l: "limitin mtopology \<sigma> l sequentially" for \<sigma> l
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
      define K where "K \<equiv> insert l (range \<sigma>)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
      interpret Submetric M d K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
      proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
        show "K \<subseteq> M"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
          unfolding K_def using l limitin_mspace \<open>S \<subseteq> M\<close> \<sigma> by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
      have "compactin mtopology K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
        unfolding K_def using \<open>S \<subseteq> M\<close> \<sigma>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
        by (force intro: compactin_sequence_with_limit [OF l])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
      then have *: "closedin sub.mtopology (K \<inter> S)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
        by (simp add: S mtopology_submetric)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
      have "\<sigma> n \<in> K \<inter> S" for n
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
        by (simp add: K_def range_subsetD \<sigma>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
      moreover have "limitin sub.mtopology \<sigma> l sequentially"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
        using l 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
        unfolding sub.limit_metric_sequentially limit_metric_sequentially
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
        by (force simp: K_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
      ultimately have "l \<in> K \<inter> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
        by (meson * \<sigma> image_subsetI sub.metric_closedin_iff_sequentially_closed) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
        by simp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
    then show "closedin mtopology S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
      unfolding metric_closedin_iff_sequentially_closed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
      using \<open>S \<subseteq> M\<close> by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
lemma (in Metric_space) k_space_mtopology: "k_space mtopology"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
  by (simp add: metrizable_imp_k_space metrizable_space_mtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
lemma k_space_euclideanreal: "k_space (euclidean :: 'a::metric_space topology)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
  using metrizable_imp_k_space metrizable_space_euclidean by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
subsection\<open>Hereditarily normal spaces\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
lemma hereditarily_B:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
  assumes "\<And>S T. separatedin X S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
      \<Longrightarrow> \<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
  shows "hereditarily normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
  unfolding hereditarily_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
  fix W
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
  assume "W \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
  show "normal_space (subtopology X W)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
    unfolding normal_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
    assume clo: "closedin (subtopology X W) S" "closedin (subtopology X W) T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
      and "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
    then have "separatedin (subtopology X W) S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
      by (simp add: separatedin_closed_sets)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
    then obtain U V where "openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
      using assms [of S T] by (meson separatedin_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
    then show "\<exists>U V. openin (subtopology X W) U \<and> openin (subtopology X W) V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
      apply (simp add: openin_subtopology_alt)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
      by (meson clo closedin_imp_subset disjnt_subset1 disjnt_subset2 inf_le2)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
lemma hereditarily_C: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
  assumes "separatedin X S T" and norm: "\<And>U. openin X U \<Longrightarrow> normal_space (subtopology X U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
  shows "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
  define ST where "ST \<equiv> X closure_of S \<inter> X closure_of T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
  have subX: "S \<subseteq> topspace X" "T \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
    by (meson \<open>separatedin X S T\<close> separation_closedin_Un_gen)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
  have sub: "S \<subseteq> topspace X - ST" "T \<subseteq> topspace X - ST"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
    unfolding ST_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
    by (metis Diff_mono Diff_triv \<open>separatedin X S T\<close> Int_lower1 Int_lower2 separatedin_def)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
  have "normal_space (subtopology X (topspace X - ST))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
    by (simp add: ST_def norm closedin_Int openin_diff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
  moreover have " disjnt (subtopology X (topspace X - ST) closure_of S)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
                         (subtopology X (topspace X - ST) closure_of T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
    using Int_absorb1 ST_def sub by (fastforce simp: disjnt_iff closure_of_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
    using sub subX
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
    apply (simp add: normal_space_closures)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
    by (metis ST_def closedin_Int closedin_closure_of closedin_def openin_trans_full)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
lemma hereditarily_normal_space: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
  "hereditarily normal_space X \<longleftrightarrow> (\<forall>U. openin X U \<longrightarrow> normal_space(subtopology X U))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
  by (metis hereditarily_B hereditarily_C hereditarily)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
lemma hereditarily_normal_separation:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
  "hereditarily normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
        (\<forall>S T. separatedin X S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
             \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
  by (metis hereditarily_B hereditarily_C hereditarily)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
lemma metrizable_imp_hereditarily_normal_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
   "metrizable_space X \<Longrightarrow> hereditarily normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
  by (simp add: hereditarily metrizable_imp_normal_space metrizable_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
lemma metrizable_space_separation:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
   "\<lbrakk>metrizable_space X; separatedin X S T\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
    \<Longrightarrow> \<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
  by (metis hereditarily hereditarily_C metrizable_imp_hereditarily_normal_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
lemma hereditarily_normal_separation_pairwise:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
   "hereditarily normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
    (\<forall>\<U>. finite \<U> \<and> (\<forall>S \<in> \<U>. S \<subseteq> topspace X) \<and> pairwise (separatedin X) \<U>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
        \<longrightarrow> (\<exists>\<F>. (\<forall>S \<in> \<U>. openin X (\<F> S) \<and> S \<subseteq> \<F> S) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
                pairwise (\<lambda>S T. disjnt (\<F> S) (\<F> T)) \<U>))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
  assume L: ?lhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
  show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
    fix \<U>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
    assume "finite \<U>" and \<U>: "\<forall>S\<in>\<U>. S \<subseteq> topspace X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
      and pw\<U>: "pairwise (separatedin X) \<U>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
    have "\<exists>V W. openin X V \<and> openin X W \<and> S \<subseteq> V \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
                    (\<forall>T. T \<in> \<U> \<and> T \<noteq> S \<longrightarrow> T \<subseteq> W) \<and> disjnt V W" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
      if "S \<in> \<U>" for S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
      have "separatedin X S (\<Union>(\<U> - {S}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
        by (metis \<U> \<open>finite \<U>\<close> pw\<U> finite_Diff pairwise_alt separatedin_Union(1) that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
      with L show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
        unfolding hereditarily_normal_separation
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
        by (smt (verit) Diff_iff UnionI empty_iff insert_iff subset_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
    then obtain \<F> \<G> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
        where *: "\<And>S. S \<in> \<U> \<Longrightarrow> S \<subseteq> \<F> S \<and> (\<forall>T. T \<in> \<U> \<and> T \<noteq> S \<longrightarrow> T \<subseteq> \<G> S)" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
        and ope: "\<And>S. S \<in> \<U> \<Longrightarrow> openin X (\<F> S) \<and> openin X (\<G> S)" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
        and dis: "\<And>S. S \<in> \<U> \<Longrightarrow> disjnt (\<F> S) (\<G> S)" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
    define \<H> where "\<H> \<equiv> \<lambda>S. \<F> S \<inter> (\<Inter>T \<in> \<U> - {S}. \<G> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
    show "\<exists>\<F>. (\<forall>S\<in>\<U>. openin X (\<F> S) \<and> S \<subseteq> \<F> S) \<and> pairwise (\<lambda>S T. disjnt (\<F> S) (\<F> T)) \<U>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
    proof (intro exI conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
      show "openin X (\<H> S)" if "S \<in> \<U>" for S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
        unfolding \<H>_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
        by (smt (verit) ope that DiffD1 \<open>finite \<U>\<close> finite_Diff finite_imageI imageE openin_Int_Inter)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
      show "S \<subseteq> \<H> S" if "S \<in> \<U>" for S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
        unfolding \<H>_def using "*" that by auto 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
    show "pairwise (\<lambda>S T. disjnt (\<H> S) (\<H> T)) \<U>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
      using dis by (fastforce simp: disjnt_iff pairwise_alt \<H>_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
  assume R: ?rhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
    unfolding hereditarily_normal_separation
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
    assume "separatedin X S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
    show "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
    proof (cases "T=S")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
        using \<open>separatedin X S T\<close> by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
      have "pairwise (separatedin X) {S, T}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
        by (simp add: \<open>separatedin X S T\<close> pairwise_insert separatedin_sym)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
      moreover have "\<forall>S\<in>{S, T}. S \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
        by (metis \<open>separatedin X S T\<close> insertE separatedin_def singletonD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
        ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
        using R by (smt (verit) False finite.emptyI finite.insertI insertCI pairwiseD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
lemma hereditarily_normal_space_perfect_map_image:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
   "\<lbrakk>hereditarily normal_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> hereditarily normal_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
  unfolding perfect_map_def proper_map_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
  by (meson hereditarily_normal_space_continuous_closed_map_image)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
lemma regular_second_countable_imp_hereditarily_normal_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
  assumes "regular_space X \<and> second_countable X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
  shows  "hereditarily normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
  unfolding hereditarily
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
  proof (intro regular_Lindelof_imp_normal_space strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
  show "regular_space (subtopology X S)" for S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
    by (simp add: assms regular_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  show "Lindelof_space (subtopology X S)" for S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
    using assms by (simp add: second_countable_imp_Lindelof_space second_countable_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
subsection\<open>Completely regular spaces\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
definition completely_regular_space where
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
 "completely_regular_space X \<equiv>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
    \<forall>S x. closedin X S \<and> x \<in> topspace X - S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
          \<longrightarrow> (\<exists>f::'a\<Rightarrow>real. continuous_map X (top_of_set {0..1}) f \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
                  f x = 0 \<and> (f ` S \<subseteq> {1}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
lemma homeomorphic_completely_regular_space_aux:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
  assumes X: "completely_regular_space X" and hom: "X homeomorphic_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
  shows "completely_regular_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
  obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
    and fg: "(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
    using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
    unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
    fix S x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
    assume A: "closedin Y S" and x: "x \<in> topspace Y" and "x \<notin> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
    then have "closedin X (g`S)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
      using hmg homeomorphic_map_closedness_eq by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
    moreover have "g x \<notin> g`S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
      by (meson A x \<open>x \<notin> S\<close> closedin_subset hmg homeomorphic_imp_injective_map inj_on_image_mem_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
    ultimately obtain \<phi> where \<phi>: "continuous_map X (top_of_set {0..1::real}) \<phi> \<and> \<phi> (g x) = 0 \<and> \<phi> ` g`S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
      by (metis DiffI X completely_regular_space_def hmg homeomorphic_imp_surjective_map image_eqI x)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
    then have "continuous_map Y (top_of_set {0..1::real}) (\<phi> \<circ> g)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
      by (meson continuous_map_compose hmg homeomorphic_imp_continuous_map)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
    then show "\<exists>\<psi>. continuous_map Y (top_of_set {0..1::real}) \<psi> \<and> \<psi> x = 0 \<and> \<psi> ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
      by (metis \<phi> comp_apply image_comp)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
lemma homeomorphic_completely_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
  assumes "X homeomorphic_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  shows "completely_regular_space X \<longleftrightarrow> completely_regular_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
  by (meson assms homeomorphic_completely_regular_space_aux homeomorphic_space_sym)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
lemma completely_regular_space_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
   "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
     (\<forall>S x. closedin X S \<longrightarrow> x \<in> topspace X - S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
           \<longrightarrow> (\<exists>f. continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
  have "\<exists>f. continuous_map X (top_of_set {0..1::real}) f \<and> f x = 0 \<and> f ` S \<subseteq> {1}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
    if "closedin X S" "x \<in> topspace X - S" and f: "continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
    show "continuous_map X (top_of_set {0..1}) (\<lambda>x. max 0 (min (f x) 1))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
      using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
      by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
  qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
  with continuous_map_in_subtopology show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
    unfolding completely_regular_space_def by metis 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
text \<open>As above, but with @{term openin}\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
lemma completely_regular_space_alt':
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
   "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
     (\<forall>S x. openin X S \<longrightarrow> x \<in> S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
           \<longrightarrow> (\<exists>f. continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` (topspace X - S) \<subseteq> {1}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
  apply (simp add: completely_regular_space_alt all_closedin)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
  by (meson openin_subset subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
lemma completely_regular_space_gen_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
  assumes "a \<noteq> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
  shows "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
         (\<forall>S x. closedin X S \<longrightarrow> x \<in> topspace X - S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
               \<longrightarrow> (\<exists>f. continuous_map X euclidean f \<and> f x = a \<and> (f ` S \<subseteq> {b})))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
  have "\<exists>f. continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
    if "closedin X S" "x \<in> topspace X - S" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
        and f: "continuous_map X euclidean f \<and> f x = a \<and> f ` S \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
    show "continuous_map X euclideanreal ((\<lambda>x. inverse(b - a) * (x - a)) \<circ> f)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
      using that by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
  qed (use that assms in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
  moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
  have "\<exists>f. continuous_map X euclidean f \<and> f x = a \<and> f ` S \<subseteq> {b}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
    if "closedin X S" "x \<in> topspace X - S" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
        and f: "continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
    show "continuous_map X euclideanreal ((\<lambda>x. a + (b - a) * x) \<circ> f)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
      using that by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
  qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
    unfolding completely_regular_space_alt by meson
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
text \<open>As above, but with @{term openin}\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
lemma completely_regular_space_gen_alt':
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
  assumes "a \<noteq> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
  shows "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
         (\<forall>S x. openin X S \<longrightarrow> x \<in> S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
               \<longrightarrow> (\<exists>f. continuous_map X euclidean f \<and> f x = a \<and> (f ` (topspace X - S) \<subseteq> {b})))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
  apply (simp add: completely_regular_space_gen_alt[OF assms] all_closedin)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
  by (meson openin_subset subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
lemma completely_regular_space_gen:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
  assumes "a < b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
  shows "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
         (\<forall>S x. closedin X S \<and> x \<in> topspace X - S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
               \<longrightarrow> (\<exists>f. continuous_map X (top_of_set {a..b}) f \<and> f x = a \<and> f ` S \<subseteq> {b}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
  have "\<exists>f. continuous_map X (top_of_set {a..b}) f \<and> f x = a \<and> f ` S \<subseteq> {b}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
    if "closedin X S" "x \<in> topspace X - S" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
      and f: "continuous_map X euclidean f \<and> f x = a \<and> f ` S \<subseteq> {b}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
    show "continuous_map X (top_of_set {a..b}) (\<lambda>x. max a (min (f x) b))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
      using that assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
      by (auto simp: continuous_map_in_subtopology intro!: continuous_map_real_max continuous_map_real_min)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
  qed (use that assms in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
  with continuous_map_in_subtopology assms show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
    using completely_regular_space_gen_alt [of a b]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
    by (smt (verit) atLeastAtMost_singleton atLeastatMost_empty singletonI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
lemma normal_imp_completely_regular_space_A:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
  assumes "normal_space X" "t1_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
  unfolding completely_regular_space_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
  fix x S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
  assume A: "closedin X S" "x \<notin> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
  assume "x \<in> topspace X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
  then have "closedin X {x}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
    by (simp add: \<open>t1_space X\<close> closedin_t1_singleton)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
  with A \<open>normal_space X\<close> have "\<exists>f. continuous_map X euclideanreal f \<and> f ` {x} \<subseteq> {0} \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
    using assms unfolding normal_space_iff_Urysohn_alt disjnt_iff by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
  then show "\<exists>f. continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
lemma normal_imp_completely_regular_space_B:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
  assumes "normal_space X" "regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
  unfolding completely_regular_space_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
  fix x S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
  assume "closedin X S" "x \<notin> S" "x \<in> topspace X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
  then obtain U C where "openin X U" "closedin X C" "x \<in> U" "U \<subseteq> C" "C \<subseteq> topspace X - S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
    unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of closedin_def by (metis Diff_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
  then obtain f where "continuous_map X euclideanreal f \<and> f ` C \<subseteq> {0} \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
    using assms unfolding normal_space_iff_Urysohn_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
    by (metis Diff_iff \<open>closedin X S\<close> disjnt_iff subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
  then show "\<exists>f. continuous_map X euclideanreal f \<and> f x = 0 \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
    by (meson \<open>U \<subseteq> C\<close> \<open>x \<in> U\<close> image_subset_iff singletonD subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
lemma normal_imp_completely_regular_space_gen:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
   "\<lbrakk>normal_space X; t1_space X \<or> Hausdorff_space X \<or> regular_space X\<rbrakk> \<Longrightarrow> completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
  using normal_imp_completely_regular_space_A normal_imp_completely_regular_space_B t1_or_Hausdorff_space by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
lemma normal_imp_completely_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
   "\<lbrakk>normal_space X; Hausdorff_space X \<or> regular_space X\<rbrakk> \<Longrightarrow> completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
  by (simp add: normal_imp_completely_regular_space_gen)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
lemma (in Metric_space) completely_regular_space_mtopology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
   "completely_regular_space mtopology"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
  by (simp add: normal_imp_completely_regular_space normal_space_mtopology regular_space_mtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
lemma metrizable_imp_completely_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
   "metrizable_space X \<Longrightarrow> completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
  by (simp add: metrizable_imp_normal_space metrizable_imp_regular_space normal_imp_completely_regular_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
lemma completely_regular_space_discrete_topology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
   "completely_regular_space(discrete_topology U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
  by (simp add: normal_imp_completely_regular_space normal_space_discrete_topology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
lemma completely_regular_space_subtopology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
  assumes "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
  shows "completely_regular_space (subtopology X S)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
  unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
  fix A x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
  assume "closedin (subtopology X S) A" and x: "x \<in> topspace (subtopology X S)" and "x \<notin> A"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
  then obtain T where "closedin X T" "A = S \<inter> T" "x \<in> topspace X" "x \<in> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
    by (force simp: closedin_subtopology_alt image_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
  then show "\<exists>f. continuous_map (subtopology X S) (top_of_set {0::real..1}) f \<and> f x = 0 \<and> f ` A \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
    using assms \<open>x \<notin> A\<close>  
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
    apply (simp add: completely_regular_space_def continuous_map_from_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
    using continuous_map_from_subtopology by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
lemma completely_regular_space_retraction_map_image:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
   " \<lbrakk>retraction_map X Y r; completely_regular_space X\<rbrakk> \<Longrightarrow> completely_regular_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
  using completely_regular_space_subtopology hereditary_imp_retractive_property homeomorphic_completely_regular_space by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
lemma completely_regular_imp_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
  assumes "completely_regular_space X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
  shows "regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
  have *: "\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
    if contf: "continuous_map X euclideanreal f" and a: "a \<in> topspace X - C" and "closedin X C"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
      and fim: "f ` topspace X \<subseteq> {0..1}" and f0: "f a = 0" and f1: "f ` C \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
    for C a f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
    show "openin X {x \<in> topspace X. f x \<in> {..<1 / 2}}" "openin X {x \<in> topspace X. f x \<in> {1 / 2<..}}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
      using openin_continuous_map_preimage [OF contf]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
      by (meson open_lessThan open_greaterThan open_openin)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
    show "a \<in> {x \<in> topspace X. f x \<in> {..<1 / 2}}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
      using a f0 by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
    show "C \<subseteq> {x \<in> topspace X. f x \<in> {1 / 2<..}}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
      using \<open>closedin X C\<close> f1 closedin_subset by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
  qed (auto simp: disjnt_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
    unfolding completely_regular_space_def regular_space_def continuous_map_in_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
    by (meson "*")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1087
proposition locally_compact_regular_imp_completely_regular_space:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
  assumes "locally_compact_space X" "Hausdorff_space X \<or> regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
  unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
  fix S x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
  assume "closedin X S" and "x \<in> topspace X" and "x \<notin> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
  have "regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
    using assms locally_compact_Hausdorff_imp_regular_space by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
  then have nbase: "neighbourhood_base_of (\<lambda>C. compactin X C \<and> closedin X C) X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
    using assms(1) locally_compact_regular_space_neighbourhood_base by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
  then obtain U M where "openin X U" "compactin X M" "closedin X M" "x \<in> U" "U \<subseteq> M" "M \<subseteq> topspace X - S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
    unfolding neighbourhood_base_of by (metis (no_types, lifting) Diff_iff \<open>closedin X S\<close> \<open>x \<in> topspace X\<close> \<open>x \<notin> S\<close> closedin_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
  then have "M \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
    by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
  obtain V K where "openin X V" "closedin X K" "x \<in> V" "V \<subseteq> K" "K \<subseteq> U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
    by (metis (no_types, lifting) \<open>openin X U\<close> \<open>x \<in> U\<close> neighbourhood_base_of nbase)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
  have "compact_space (subtopology X M)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
    by (simp add: \<open>compactin X M\<close> compact_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
  then have "normal_space (subtopology X M)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
    by (simp add: \<open>regular_space X\<close> compact_Hausdorff_or_regular_imp_normal_space regular_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
  moreover have "closedin (subtopology X M) K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
    using \<open>K \<subseteq> U\<close> \<open>U \<subseteq> M\<close> \<open>closedin X K\<close> closedin_subset_topspace by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
  moreover have "closedin (subtopology X M) (M - U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
    by (simp add: \<open>closedin X M\<close> \<open>openin X U\<close> closedin_diff closedin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
  moreover have "disjnt K (M - U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
    by (meson DiffD2 \<open>K \<subseteq> U\<close> disjnt_iff subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
  ultimately obtain f::"'a\<Rightarrow>real" where contf: "continuous_map (subtopology X M) (top_of_set {0..1}) f" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
    and f0: "f ` K \<subseteq> {0}" and f1: "f ` (M - U) \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
    using Urysohn_lemma [of "subtopology X M" K "M-U" 0 1] by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
  then obtain g::"'a\<Rightarrow>real" where contg: "continuous_map (subtopology X M) euclidean g" and gim: "g ` M \<subseteq> {0..1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
    and g0: "\<And>x. x \<in> K \<Longrightarrow> g x = 0" and g1: "\<And>x. \<lbrakk>x \<in> M; x \<notin> U\<rbrakk> \<Longrightarrow> g x = 1"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1119
    using \<open>M \<subseteq> topspace X\<close> by (force simp: continuous_map_in_subtopology image_subset_iff)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
  show "\<exists>f::'a\<Rightarrow>real. continuous_map X (top_of_set {0..1}) f \<and> f x = 0 \<and> f ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
    show "continuous_map X (top_of_set {0..1}) (\<lambda>x. if x \<in> M then g x else 1)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
      unfolding continuous_map_closedin
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
    proof (intro strip conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
      fix C
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
      assume C: "closedin (top_of_set {0::real..1}) C"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
      have eq: "{x \<in> topspace X. (if x \<in> M then g x else 1) \<in> C} = {x \<in> M. g x \<in> C} \<union> (if 1 \<in> C then topspace X - U else {})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
        using \<open>U \<subseteq> M\<close> \<open>M \<subseteq> topspace X\<close> g1 by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
      show "closedin X {x \<in> topspace X. (if x \<in> M then g x else 1) \<in> C}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
        unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
      proof (intro closedin_Un)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
        have "closedin euclidean C"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
          using C closed_closedin closedin_closed_trans by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
        then have "closedin (subtopology X M) {x \<in> M. g x \<in> C}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
          using closedin_continuous_map_preimage_gen [OF contg] \<open>M \<subseteq> topspace X\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
        then show "closedin X {x \<in> M. g x \<in> C}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
          using \<open>closedin X M\<close> closedin_trans_full by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
      qed (use \<open>openin X U\<close> in force)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
    qed (use gim in force)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
    show "(if x \<in> M then g x else 1) = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
      using \<open>U \<subseteq> M\<close> \<open>V \<subseteq> K\<close> g0 \<open>x \<in> U\<close> \<open>x \<in> V\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
    show "(\<lambda>x. if x \<in> M then g x else 1) ` S \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
      using \<open>M \<subseteq> topspace X - S\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
lemma completely_regular_eq_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
   "locally_compact_space X
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
        \<Longrightarrow> (completely_regular_space X \<longleftrightarrow> regular_space X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
  using completely_regular_imp_regular_space locally_compact_regular_imp_completely_regular_space 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
  by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
lemma completely_regular_space_prod_topology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
   "completely_regular_space (prod_topology X Y) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1155
      (prod_topology X Y) = trivial_topology \<or>
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
      completely_regular_space X \<and> completely_regular_space Y" (is "?lhs=?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
    by (rule topological_property_of_prod_component) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
       (auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
  show ?lhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1164
  proof (cases "(prod_topology X Y) = trivial_topology")
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
    then have X: "completely_regular_space X" and Y: "completely_regular_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
      using R by blast+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
      unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
    proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
      fix W x y
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
      assume "openin (prod_topology X Y) W" and "(x, y) \<in> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
      then obtain U V where "openin X U" "openin Y V" "x \<in> U" "y \<in> V" "U\<times>V \<subseteq> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
        by (force simp: openin_prod_topology_alt)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
      then have "x \<in> topspace X" "y \<in> topspace Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
        using openin_subset by fastforce+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
      obtain f where contf: "continuous_map X euclideanreal f" and "f x = 0" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
        and f1: "\<And>x. x \<in> topspace X \<Longrightarrow> x \<notin> U \<Longrightarrow> f x = 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
        using X \<open>openin X U\<close> \<open>x \<in> U\<close> unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
        by (smt (verit, best) Diff_iff image_subset_iff singletonD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
      obtain g where contg: "continuous_map Y euclideanreal g" and "g y = 0" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
        and g1: "\<And>y. y \<in> topspace Y \<Longrightarrow> y \<notin> V \<Longrightarrow> g y = 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
        using Y \<open>openin Y V\<close> \<open>y \<in> V\<close> unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
        by (smt (verit, best) Diff_iff image_subset_iff singletonD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
      define h where "h \<equiv> \<lambda>(x,y). 1 - (1 - f x) * (1 - g y)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
      show "\<exists>h. continuous_map (prod_topology X Y) euclideanreal h \<and> h (x,y) = 0 \<and> h ` (topspace (prod_topology X Y) - W) \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
      proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
        have "continuous_map (prod_topology X Y) euclideanreal (f \<circ> fst)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
          using contf continuous_map_of_fst by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
        moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
        have "continuous_map (prod_topology X Y) euclideanreal (g \<circ> snd)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
          using contg continuous_map_of_snd by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
        ultimately
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
        show "continuous_map (prod_topology X Y) euclideanreal h"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
          unfolding o_def h_def case_prod_unfold
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
          by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
        show "h (x, y) = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
          by (simp add: h_def \<open>f x = 0\<close> \<open>g y = 0\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
        show "h ` (topspace (prod_topology X Y) - W) \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
          using \<open>U \<times> V \<subseteq> W\<close> f1 g1 by (force simp: h_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
  qed (force simp: completely_regular_space_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1207
proposition completely_regular_space_product_topology:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
   "completely_regular_space (product_topology X I) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1209
    (\<exists>i\<in>I. X i = trivial_topology) \<or> (\<forall>i \<in> I. completely_regular_space (X i))" 
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
    by (rule topological_property_of_product_component) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
       (auto simp: completely_regular_space_subtopology homeomorphic_completely_regular_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
  show ?lhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  1218
  proof (cases "\<exists>i\<in>I. X i = trivial_topology")
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
      unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
    proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
      fix W x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
      assume W: "openin (product_topology X I) W" and "x \<in> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
      then obtain U where finU: "finite {i \<in> I. U i \<noteq> topspace (X i)}" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
             and ope: "\<And>i. i\<in>I \<Longrightarrow> openin (X i) (U i)" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
             and x: "x \<in> Pi\<^sub>E I U" and "Pi\<^sub>E I U \<subseteq> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
        by (auto simp: openin_product_topology_alt)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
      have "\<forall>i \<in> I. openin (X i) (U i) \<and> x i \<in> U i
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
              \<longrightarrow> (\<exists>f. continuous_map (X i) euclideanreal f \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
                       f (x i) = 0 \<and> (\<forall>x \<in> topspace(X i). x \<notin> U i \<longrightarrow> f x = 1))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
        using R unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
        by (smt (verit) DiffI False image_subset_iff singletonD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
      with ope x have "\<And>i. \<exists>f. i \<in> I \<longrightarrow> continuous_map (X i) euclideanreal f \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
              f (x i) = 0 \<and> (\<forall>x \<in> topspace (X i). x \<notin> U i \<longrightarrow> f x = 1)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
      then obtain f where f: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (X i) euclideanreal (f i) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
                                             f i (x i) = 0 \<and> (\<forall>x \<in> topspace (X i). x \<notin> U i \<longrightarrow> f i x = 1)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
        by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
      define h where "h \<equiv> \<lambda>z. 1 - prod (\<lambda>i. 1 - f i (z i)) {i\<in>I. U i \<noteq> topspace(X i)}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
      show "\<exists>h. continuous_map (product_topology X I) euclideanreal h \<and> h x = 0 \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
                     h ` (topspace (product_topology X I) - W) \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
      proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
        have "continuous_map (product_topology X I) euclidean (f i \<circ> (\<lambda>x. x i))" if "i\<in>I" for i
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
          using f that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
          by (blast intro: continuous_intros continuous_map_product_projection)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
        then show "continuous_map (product_topology X I) euclideanreal h"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
          unfolding h_def o_def by (intro continuous_intros) (auto simp: finU)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
        show "h x = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
          by (simp add: h_def f)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
        show "h ` (topspace (product_topology X I) - W) \<subseteq> {1}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
          have "\<exists>i. i \<in> I \<and> U i \<noteq> topspace (X i) \<and> f i (x' i) = 1"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
            if "x' \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" "x' \<notin> W" for x'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
            using that \<open>Pi\<^sub>E I U \<subseteq> W\<close> by (smt (verit, best) PiE_iff f in_mono)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
            by (auto simp: f h_def finU)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
    qed      
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
  qed (force simp: completely_regular_space_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1265
lemma zero_dimensional_imp_completely_regular_space:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1266
  assumes "X dim_le 0" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1267
  shows "completely_regular_space X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1268
proof (clarsimp simp: completely_regular_space_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1269
  fix C a
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1270
  assume "closedin X C" and "a \<in> topspace X" and "a \<notin> C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1271
  then obtain U where "closedin X U" "openin X U" "a \<in> U" and U: "U \<subseteq> topspace X - C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1272
    using assms by (force simp add: closedin_def dimension_le_0_neighbourhood_base_of_clopen open_neighbourhood_base_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1273
  show "\<exists>f. continuous_map X (top_of_set {0::real..1}) f \<and> f a = 0 \<and> f ` C \<subseteq> {1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1274
  proof (intro exI conjI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1275
    have "continuous_map X euclideanreal (\<lambda>x. if x \<in> U then 0 else 1)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1276
      using \<open>closedin X U\<close> \<open>openin X U\<close> apply (clarsimp simp: continuous_map_def closedin_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1277
      by (smt (verit) Diff_iff mem_Collect_eq openin_subopen subset_eq)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1278
    then show "continuous_map X (top_of_set {0::real..1}) (\<lambda>x. if x \<in> U then 0 else 1)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1279
      by (auto simp: continuous_map_in_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1280
  qed (use U \<open>a \<in> U\<close> in auto)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1281
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1282
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1283
lemma zero_dimensional_imp_regular_space: "X dim_le 0 \<Longrightarrow> regular_space X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1284
  by (simp add: completely_regular_imp_regular_space zero_dimensional_imp_completely_regular_space)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1285
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
lemma (in Metric_space) t1_space_mtopology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
   "t1_space mtopology"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
  using Hausdorff_space_mtopology t1_or_Hausdorff_space by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
subsection\<open>More generally, the k-ification functor\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
definition kification_open 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
  where "kification_open \<equiv> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
          \<lambda>X S. S \<subseteq> topspace X \<and> (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
definition kification 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
  where "kification X \<equiv> topology (kification_open X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
lemma istopology_kification_open: "istopology (kification_open X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
  unfolding istopology_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
  show "kification_open X (S \<inter> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
    if "kification_open X S" and "kification_open X T" for S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
    using that unfolding kification_open_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
    by (smt (verit, best) inf.idem inf_commute inf_left_commute le_infI2 openin_Int)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
  show "kification_open X (\<Union> \<K>)" if "\<forall>K\<in>\<K>. kification_open X K" for \<K>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
    using that unfolding kification_open_def Int_Union by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
lemma openin_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
   "openin (kification X) U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
        U \<subseteq> topspace X \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
        (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> U))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
  by (metis topology_inverse' kification_def istopology_kification_open kification_open_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
lemma openin_kification_finer:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
   "openin X S \<Longrightarrow> openin (kification X) S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
  by (simp add: openin_kification openin_subset openin_subtopology_Int2)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
lemma topspace_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
   "topspace(kification X) = topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
  by (meson openin_kification openin_kification_finer openin_topspace subset_antisym)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
lemma closedin_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
   "closedin (kification X) U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
      U \<subseteq> topspace X \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
      (\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> U))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
proof (cases "U \<subseteq> topspace X")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
    by (simp add: closedin_def Diff_Int_distrib inf_commute le_infI2 openin_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
qed (simp add: closedin_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
lemma closedin_kification_finer: "closedin X S \<Longrightarrow> closedin (kification X) S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
  by (simp add: closedin_def openin_kification_finer)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
lemma kification_eq_self: "kification X = X \<longleftrightarrow> k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
  unfolding fun_eq_iff openin_kification k_space_alt openin_inject [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
  by (metis openin_closedin_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
lemma compactin_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
   "compactin (kification X) K \<longleftrightarrow> compactin X K" (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
    by (simp add: compactin_contractive openin_kification_finer)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
    unfolding compactin_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
    show "K \<subseteq> topspace (kification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
      by (simp add: R compactin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
    fix \<U>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
    assume \<U>: "Ball \<U> (openin (kification X)) \<and> K \<subseteq> \<Union> \<U>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
    then have *: "\<And>U. U \<in> \<U> \<Longrightarrow> U \<subseteq> topspace X \<and> openin (subtopology X K) (K \<inter> U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
      by (simp add: R openin_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
    have "K \<subseteq> topspace X" "compact_space (subtopology X K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
      using R compactin_subspace by force+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
    then have "\<exists>V. finite V \<and> V \<subseteq> (\<lambda>U. K \<inter> U) ` \<U> \<and> \<Union> V = topspace (subtopology X K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
      unfolding compact_space
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
      by (smt (verit, del_insts) Int_Union \<U> * image_iff inf.order_iff inf_commute topspace_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
    then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> K \<subseteq> \<Union> \<F>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
      by (metis Int_Union \<open>K \<subseteq> topspace X\<close> finite_subset_image inf.orderI topspace_subtopology_subset)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
lemma compact_space_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
   "compact_space(kification X) \<longleftrightarrow> compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
  by (simp add: compact_space_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
lemma kification_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
   "kification(kification X) = kification X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
  unfolding openin_inject [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
  show "openin (kification (kification X)) U = openin (kification X) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
  proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
    show "openin (kification (kification X)) U \<Longrightarrow> openin (kification X) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
      by (metis compactin_kification inf_commute openin_kification openin_subtopology topspace_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
  qed (simp add: openin_kification_finer)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1382
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1383
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
lemma k_space_kification [iff]: "k_space(kification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
  using kification_eq_self by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
lemma continuous_map_into_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
  assumes "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
  shows "continuous_map X (kification Y) f \<longleftrightarrow> continuous_map X Y f" (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
    by (simp add: continuous_map_def openin_kification_finer)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
  have "openin X {x \<in> topspace X. f x \<in> V}" if V: "openin (kification Y) V" for V
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
    have "openin (subtopology X K) (K \<inter> {x \<in> topspace X. f x \<in> V})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
      if "compactin X K" for K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
      have "compactin Y (f ` K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
        using R image_compactin that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
      then have "openin (subtopology Y (f ` K)) (f ` K \<inter> V)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
        by (meson V openin_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
      then obtain U where U: "openin Y U" "f`K \<inter> V = U \<inter> f`K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
        by (meson openin_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
        unfolding openin_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
      proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
        show "openin X {x \<in> topspace X. f x \<in> U}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
          using R U openin_continuous_map_preimage_gen by (simp add: continuous_map_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
      qed (use U in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
      by (metis (full_types) Collect_subset assms k_space_open)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
  with R show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
    by (simp add: continuous_map_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
lemma continuous_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
   "continuous_map X Y f \<Longrightarrow> continuous_map (kification X) Y f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
  by (simp add: continuous_map_openin_preimage_eq openin_kification_finer)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
lemma continuous_map_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
   "continuous_map X Y f \<Longrightarrow> continuous_map (kification X) (kification Y) f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
  by (simp add: continuous_map_from_kification continuous_map_into_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
lemma subtopology_kification_compact:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
  assumes "compactin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
  shows "subtopology (kification X) K = subtopology X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
  unfolding openin_inject [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
  show "openin (subtopology (kification X) K) U = openin (subtopology X K) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
    by (metis assms inf_commute openin_kification openin_subset openin_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
lemma subtopology_kification_finer:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
  assumes "openin (subtopology (kification X) S) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
  shows "openin (kification (subtopology X S)) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
  using assms 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
  by (fastforce simp: openin_subtopology_alt image_iff openin_kification subtopology_subtopology compactin_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
lemma proper_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
  assumes "k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
  shows "proper_map (kification X) Y f \<longleftrightarrow> proper_map X Y f"   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
    by (simp add: closed_map_def closedin_kification_finer proper_map_alt)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
  have "compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}" for K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
    using R proper_map_alt by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
  with R show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
    by (simp add: assms proper_map_into_k_space_eq subtopology_kification_compact)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
lemma perfect_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
   "\<lbrakk>k_space Y; perfect_map X Y f\<rbrakk> \<Longrightarrow> perfect_map(kification X) Y f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
  by (simp add: continuous_map_from_kification perfect_map_def proper_map_from_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1463
lemma k_space_perfect_map_image_eq:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
  assumes "Hausdorff_space X" "perfect_map X Y f"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
  shows "k_space X \<longleftrightarrow> k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
  show "k_space X \<Longrightarrow> k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
    using k_space_perfect_map_image assms by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
  assume "k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
  have "homeomorphic_map (kification X) X id"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
    unfolding homeomorphic_eq_injective_perfect_map
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
    proof (intro conjI perfect_map_from_composition_right [where f = id])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
  show "perfect_map (kification X) Y (f \<circ> id)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
    by (simp add: \<open>k_space Y\<close> assms(2) perfect_map_from_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
  show "continuous_map (kification X) X id"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
    by (simp add: continuous_map_from_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
qed (use assms perfect_map_def in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
  then show "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
    using homeomorphic_k_space homeomorphic_space by blast 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
subsection\<open>One-point compactifications and the Alexandroff extension construction\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
lemma one_point_compactification_dense:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
   "\<lbrakk>compact_space X; \<not> compactin X (topspace X - {a})\<rbrakk> \<Longrightarrow> X closure_of (topspace X - {a}) = topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
  unfolding closure_of_complement
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
  by (metis Diff_empty closedin_compact_space interior_of_eq_empty openin_closedin_eq subset_singletonD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
lemma one_point_compactification_interior:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
   "\<lbrakk>compact_space X; \<not> compactin X (topspace X - {a})\<rbrakk> \<Longrightarrow> X interior_of {a} = {}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
  by (simp add: interior_of_eq_empty_complement one_point_compactification_dense)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1494
proposition kc_space_one_point_compactification_gen:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
  shows "kc_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
         openin X (topspace X - {a}) \<and> (\<forall>K. compactin X K \<and> a\<notin>K \<longrightarrow> closedin X K) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
         k_space (subtopology X (topspace X - {a})) \<and> kc_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
 (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
  assume L: ?lhs show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
    show "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
      using L kc_imp_t1_space t1_space_openin_delete_alt by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
    then show "k_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
      by (simp add: L assms k_space_open_subtopology_aux)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
    show "closedin X k" if "compactin X k \<and> a \<notin> k" for k :: "'a set"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
      using L kc_space_def that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
    show "kc_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
      by (simp add: L kc_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
    unfolding kc_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
    fix S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
    assume "compactin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
    then have "S \<subseteq>topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
      by (simp add: compactin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
    show "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
    proof (cases "a \<in> S")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
      then have "topspace X - S = topspace X - {a} - (S - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
      moreover have "openin X (topspace X - {a} - (S - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
      proof (rule openin_trans_full)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
        show "openin (subtopology X (topspace X - {a})) (topspace X - {a} - (S - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
        proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
          show "openin (subtopology X (topspace X - {a})) (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
            using R openin_open_subtopology by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
          have "closedin (subtopology X ((topspace X - {a}) \<inter> K)) (K \<inter> (S - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
            if "compactin X K" "K \<subseteq> topspace X - {a}" for K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
          proof (intro closedin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
            show "closedin X (K \<inter> (S - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
              using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
              by (metis IntD1 Int_insert_right_if0 R True \<open>compactin X S\<close> closed_Int_compactin insert_Diff subset_Diff_insert)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
          qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
          moreover have "k_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
            using R by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
          moreover have "S-{a} \<subseteq> topspace X \<and> S-{a} \<subseteq> topspace X - {a}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
            using \<open>S \<subseteq> topspace X\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
          ultimately show "closedin (subtopology X (topspace X - {a})) (S - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
            using \<open>S \<subseteq> topspace X\<close> True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
            by (simp add: k_space_def compactin_subtopology subtopology_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
        qed 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
        show "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
          by (simp add: R)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
      ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
        by (simp add: \<open>S \<subseteq> topspace X\<close> closedin_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
        by (simp add: R \<open>compactin X S\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
  
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
inductive Alexandroff_open for X where
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
  base: "openin X U \<Longrightarrow> Alexandroff_open X (Some ` U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
| ext: "\<lbrakk>compactin X C; closedin X C\<rbrakk> \<Longrightarrow> Alexandroff_open X (insert None (Some ` (topspace X - C)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
78128
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1565
hide_fact (open) base ext
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1566
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
lemma Alexandroff_open_iff: "Alexandroff_open X S \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
   (\<exists>U. (S = Some ` U \<and> openin X U) \<or> (S = insert None (Some ` (topspace X - U)) \<and> compactin X U \<and> closedin X U))"
78128
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1569
  by (meson Alexandroff_open.cases Alexandroff_open.ext Alexandroff_open.base)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1570
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
lemma Alexandroff_open_Un_aux:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
  assumes U: "openin X U" and "Alexandroff_open X T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
  shows  "Alexandroff_open X (Some ` U \<union> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
  using \<open>Alexandroff_open X T\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
proof (induction rule: Alexandroff_open.induct)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
  case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
    by (metis Alexandroff_open.base U image_Un openin_Un)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1580
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
  have "U \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
    by (simp add: U openin_subset)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
  then have eq: "Some ` U \<union> insert None (Some ` (topspace X - C)) = insert None (Some ` (topspace X - (C \<inter> (topspace X - U))))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
    by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
  have "closedin X (C \<inter> (topspace X - U))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
    using U ext.hyps(2) by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
  moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
  have "compactin X (C \<inter> (topspace X - U))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
    using U compact_Int_closedin ext.hyps(1) by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
  ultimately show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
    unfolding eq using Alexandroff_open.ext by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
lemma Alexandroff_open_Un:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
  assumes "Alexandroff_open X S" and "Alexandroff_open X T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
  shows "Alexandroff_open X (S \<union> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
  using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1598
proof (induction rule: Alexandroff_open.induct)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
  case (base U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
    by (simp add: Alexandroff_open_Un_aux)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
    by (smt (verit, best) Alexandroff_open_Un_aux Alexandroff_open_iff Un_commute Un_insert_left closedin_def insert_absorb2)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
lemma Alexandroff_open_Int_aux:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
  assumes U: "openin X U" and "Alexandroff_open X T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
  shows  "Alexandroff_open X (Some ` U \<inter> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
  using \<open>Alexandroff_open X T\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
proof (induction rule: Alexandroff_open.induct)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
  case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
    by (metis Alexandroff_open.base U image_Int inj_Some openin_Int)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
  have eq: "Some ` U \<inter> insert None (Some ` (topspace X - C)) = Some ` (topspace X - (C \<union> (topspace X - U)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
    by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
  have "openin X (topspace X - (C \<union> (topspace X - U)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
    using U ext.hyps(2) by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
    unfolding eq using Alexandroff_open.base by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1626
proposition istopology_Alexandroff_open: "istopology (Alexandroff_open X)"
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
  unfolding istopology_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
  fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
  assume "Alexandroff_open X S" and "Alexandroff_open X T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
  then show "Alexandroff_open X (S \<inter> T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
  proof (induction rule: Alexandroff_open.induct)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
    case (base U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
    then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
      using Alexandroff_open_Int_aux by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
    case EC: (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
    show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
      using \<open>Alexandroff_open X T\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
    proof (induction rule: Alexandroff_open.induct)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
      case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
      then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
        by (metis Alexandroff_open.ext Alexandroff_open_Int_aux EC.hyps inf_commute)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
      case (ext D)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
      have eq: "insert None (Some ` (topspace X - C)) \<inter> insert None (Some ` (topspace X - D))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
              = insert None (Some ` (topspace X - (C \<union> D)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
      show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
        unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
        by (simp add: Alexandroff_open.ext EC.hyps closedin_Un compactin_Un ext.hyps)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
  fix \<K>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
  assume \<section>: "\<forall>K\<in>\<K>. Alexandroff_open X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
  show "Alexandroff_open X (\<Union>\<K>)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
  proof (cases "None \<in> \<Union>\<K>")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
    case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
    have "\<forall>K\<in>\<K>. \<exists>U. (openin X U \<and> K = Some ` U) \<or> (K = insert None (Some ` (topspace X - U)) \<and> compactin X U \<and> closedin X U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
      by (metis \<section> Alexandroff_open_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
    then obtain U where U: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
      "\<And>K. K \<in> \<K> \<Longrightarrow> openin X (U K) \<and> K = Some ` (U K) 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
                    \<or> (K = insert None (Some ` (topspace X - U K)) \<and> compactin X (U K) \<and> closedin X (U K))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
    define \<K>N where "\<K>N \<equiv> {K \<in> \<K>. None \<in> K}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
    define A where "A \<equiv> \<Union>K\<in>\<K>-\<K>N. U K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
    define B where "B \<equiv> \<Inter>K\<in>\<K>N. U K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
    have U1: "\<And>K. K \<in> \<K>-\<K>N \<Longrightarrow> openin X (U K) \<and> K = Some ` (U K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
      using U \<K>N_def by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1671
    have U2: "\<And>K. K \<in> \<K>N \<Longrightarrow> K = insert None (Some ` (topspace X - U K)) \<and> compactin X (U K) \<and> closedin X (U K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
      using U \<K>N_def by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
    have eqA: "\<Union>(\<K>-\<K>N) = Some ` A"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1674
    proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
      show "\<Union> (\<K> - \<K>N) \<subseteq> Some ` A"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
        by (metis A_def Sup_le_iff U1 UN_upper subset_image_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
      show "Some ` A \<subseteq> \<Union> (\<K> - \<K>N)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
        using A_def U1 by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
    have eqB: "\<Union>\<K>N = insert None (Some ` (topspace X - B))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
      using U2 True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
      by (auto simp: B_def image_iff \<K>N_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
    have "\<Union>\<K> = \<Union>\<K>N \<union> \<Union>(\<K>-\<K>N)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
      by (auto simp: \<K>N_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
    then have eq: "\<Union>\<K> = (Some ` A) \<union> (insert None (Some ` (topspace X - B)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
      by (simp add: eqA eqB Un_commute)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
      unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
    proof (intro Alexandroff_open_Un Alexandroff_open.intros)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
      show "openin X A"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
        using A_def U1 by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
      show "closedin X B"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
        unfolding B_def using U2 True \<K>N_def by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
      show "compactin X B"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
        by (metis B_def U2 eqB Inf_lower Union_iff \<open>closedin X B\<close> closed_compactin imageI insertI1)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
    then have "\<forall>K\<in>\<K>. \<exists>U. openin X U \<and> K = Some ` U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
      by (metis Alexandroff_open.simps UnionI \<section> insertCI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
    then obtain U where U: "\<forall>K\<in>\<K>. openin X (U K) \<and> K = Some ` (U K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
    then have eq: "\<Union>\<K> = Some ` (\<Union> K\<in>\<K>. U K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
      using image_iff by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
    show ?thesis
78128
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1706
      unfolding eq by (simp add: U Alexandroff_open.base openin_clauses(3))
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
definition Alexandroff_compactification where
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
  "Alexandroff_compactification X \<equiv> topology (Alexandroff_open X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
lemma openin_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
   "openin(Alexandroff_compactification X) V \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
        (\<exists>U. openin X U \<and> V = Some ` U) \<or>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
        (\<exists>C. compactin X C \<and> closedin X C \<and> V = insert None (Some ` (topspace X - C)))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
  by (auto simp: Alexandroff_compactification_def istopology_Alexandroff_open Alexandroff_open.simps)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
lemma topspace_Alexandroff_compactification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
   "topspace(Alexandroff_compactification X) = insert None (Some ` topspace X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
   (is "?lhs = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
  show "?lhs \<subseteq> ?rhs"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1726
    by (force simp: topspace_def openin_Alexandroff_compactification)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
  have "None \<in> topspace (Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
    by (meson closedin_empty compactin_empty insert_subset openin_Alexandroff_compactification openin_subset)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
  moreover have "Some x \<in> topspace (Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
    if "x \<in> topspace X" for x 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
    by (meson that imageI openin_Alexandroff_compactification openin_subset openin_topspace subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
  ultimately show "?rhs \<subseteq> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
    by (auto simp: image_subset_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
lemma closedin_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
   "closedin (Alexandroff_compactification X) C \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
      (\<exists>K. compactin X K \<and> closedin X K \<and> C = Some ` K) \<or>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
      (\<exists>U. openin X U \<and> C = topspace(Alexandroff_compactification X) - Some ` U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
    apply (clarsimp simp: closedin_def openin_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
    by (smt (verit) Diff_Diff_Int None_notin_image_Some image_set_diff inf.absorb_iff2 inj_Some insert_Diff_if subset_insert)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
    using openin_subset 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
    by (fastforce simp: closedin_def openin_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
lemma openin_Alexandroff_compactification_image_Some [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
   "openin(Alexandroff_compactification X) (Some ` U) \<longleftrightarrow> openin X U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
  by (auto simp: openin_Alexandroff_compactification inj_image_eq_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1754
lemma closedin_Alexandroff_compactification_image_Some [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
   "closedin (Alexandroff_compactification X) (Some ` K) \<longleftrightarrow> compactin X K \<and> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
  by (auto simp: closedin_Alexandroff_compactification inj_image_eq_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
lemma open_map_Some: "open_map X (Alexandroff_compactification X) Some"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
  using open_map_def openin_Alexandroff_compactification by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
lemma continuous_map_Some: "continuous_map X (Alexandroff_compactification X) Some"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
  unfolding continuous_map_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
  assume "openin (Alexandroff_compactification X) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
  then consider V where "openin X V" "U = Some ` V" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
    | C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
    by (auto simp: openin_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
  then show "openin X {x \<in> topspace X. Some x \<in> U}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
  proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
    case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
      using openin_subopen openin_subset by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
    case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
      by (simp add: closedin_def image_iff set_diff_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
lemma embedding_map_Some: "embedding_map X (Alexandroff_compactification X) Some"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
  by (simp add: continuous_map_Some injective_open_imp_embedding_map open_map_Some)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
lemma compact_space_Alexandroff_compactification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
   "compact_space(Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
proof (clarsimp simp: compact_space_alt image_subset_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
  fix \<U> U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
  assume ope [rule_format]: "\<forall>U\<in>\<U>. openin (Alexandroff_compactification X) U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
      and cover: "\<forall>x\<in>topspace X. \<exists>X\<in>\<U>. Some x \<in> X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
      and "U \<in> \<U>" "None \<in> U"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
  then have Usub: "U \<subseteq> insert None (Some ` topspace X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
    by (metis openin_subset topspace_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
  with ope [OF \<open>U \<in> \<U>\<close>] \<open>None \<in> U\<close>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
  obtain C where C: "compactin X C \<and> closedin X C \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
          insert None (Some ` topspace X) - U = Some ` C"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
    by (auto simp: openin_closedin closedin_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
  then have D: "compactin (Alexandroff_compactification X) (insert None (Some ` topspace X) - U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
    by (metis continuous_map_Some image_compactin)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
  consider V where "openin X V" "U = Some ` V" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
    | C where "compactin X C" "closedin X C" "U = insert None (Some ` (topspace X - C))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
    using ope [OF \<open>U \<in> \<U>\<close>] by (auto simp: openin_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
  then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> (\<exists>X\<in>\<F>. None \<in> X) \<and> (\<forall>x\<in>topspace X. \<exists>X\<in>\<F>. Some x \<in> X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
  proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
    case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
      using \<open>None \<in> U\<close> by blast      
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
    case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
    obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<U>" "insert None (Some ` topspace X) - U \<subseteq> \<Union>\<F>"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
      by (smt (verit, del_insts) C D Union_iff compactinD compactin_subset_topspace cover image_subset_iff ope subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
    with \<open>U \<in> \<U>\<close> show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
      by (rule_tac x="insert U \<F>" in exI) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
lemma topspace_Alexandroff_compactification_delete:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
   "topspace(Alexandroff_compactification X) - {None} = Some ` topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
  by simp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
lemma Alexandroff_compactification_dense:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
  assumes "\<not> compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
  shows "(Alexandroff_compactification X) closure_of (Some ` topspace X) =
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
         topspace(Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
  unfolding topspace_Alexandroff_compactification_delete [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
proof (intro one_point_compactification_dense)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
  show "\<not> compactin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
    using assms compact_space_proper_map_preimage compact_space_subtopology embedding_map_Some embedding_map_def homeomorphic_imp_proper_map by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
lemma t0_space_one_point_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
  assumes "compact_space X \<and> openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
  shows "t0_space X \<longleftrightarrow> t0_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
    using t0_space_subtopology by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
    unfolding t0_space_def by (bestsimp simp flip: Int_Diff dest: openin_trans_full)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
lemma t0_space_Alexandroff_compactification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
   "t0_space (Alexandroff_compactification X) \<longleftrightarrow> t0_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
  using t0_space_one_point_compactification [of "Alexandroff_compactification X" None]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
  using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t0_space by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
lemma t1_space_one_point_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
  assumes Xa: "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
    and \<section>: "\<And>K. \<lbrakk>compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\<rbrakk> \<Longrightarrow> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1852
  shows "t1_space X \<longleftrightarrow> t1_space (subtopology X (topspace X - {a}))"  (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
    using t1_space_subtopology by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1857
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
    unfolding t1_space_closedin_singleton
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
    fix x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
    assume "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
    show "closedin X {x}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
    proof (cases "x=a")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
        using \<open>x \<in> topspace X\<close> Xa closedin_def by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
        by (simp add: "\<section>" False R \<open>x \<in> topspace X\<close> closedin_t1_singleton)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1871
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
lemma closedin_Alexandroff_I: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
  assumes "compactin (Alexandroff_compactification X) K" "K \<subseteq> Some ` topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
          "closedin (Alexandroff_compactification X) T" "K = T \<inter> Some ` topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
  shows "closedin (Alexandroff_compactification X) K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
  obtain S where S: "S \<subseteq> topspace X" "K = Some ` S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
    by (meson \<open>K \<subseteq> Some ` topspace X\<close> subset_imageE)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
  with assms have "compactin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1883
    by (metis compactin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_compactness)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
  moreover have "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1885
    using assms S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
    by (metis closedin_subtopology embedding_map_Some embedding_map_def homeomorphic_map_closedness)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
    by (simp add: S)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
lemma t1_space_Alexandroff_compactification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
  "t1_space(Alexandroff_compactification X) \<longleftrightarrow> t1_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
  have "openin (Alexandroff_compactification X) (topspace (Alexandroff_compactification X) - {None})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
    using t1_space_one_point_compactification [of "Alexandroff_compactification X" None]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
    using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_t1_space 
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1900
    by (fastforce simp: compactin_subtopology closedin_Alexandroff_I closedin_subtopology)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
lemma kc_space_one_point_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
    and ope: "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
    and \<section>: "\<And>K. \<lbrakk>compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\<rbrakk>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
                \<Longrightarrow> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
  shows "kc_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
         k_space (subtopology X (topspace X - {a})) \<and> kc_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
  have "closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
    if "kc_space (subtopology X (topspace X - {a}))" and "compactin X K" "a \<notin> K" for K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
    using that unfolding kc_space_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
    by (metis "\<section>" Diff_empty compactin_subspace compactin_subtopology subset_Diff_insert)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
    by (metis \<open>compact_space X\<close> kc_space_one_point_compactification_gen ope)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
lemma kc_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
  "kc_space(Alexandroff_compactification X) \<longleftrightarrow> (k_space X \<and> kc_space X)" (is "kc_space ?Y = _")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
  have "kc_space (Alexandroff_compactification X) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
      (k_space (subtopology ?Y (topspace ?Y - {None})) \<and> kc_space (subtopology ?Y (topspace ?Y - {None})))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
    by (rule kc_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
  also have "... \<longleftrightarrow> k_space X \<and> kc_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
    using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_k_space homeomorphic_kc_space by simp blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1929
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1932
proposition regular_space_one_point_compactification:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
  assumes "compact_space X" and ope: "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
    and \<section>: "\<And>K. \<lbrakk>compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\<rbrakk> \<Longrightarrow> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1935
  shows "regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1936
           regular_space (subtopology X (topspace X - {a})) \<and> locally_compact_space (subtopology X (topspace X - {a}))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
    (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1938
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1940
    using assms(1) compact_imp_locally_compact_space locally_compact_space_open_subset ope regular_space_subtopology by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
  let ?Xa = "subtopology X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1944
    unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of imp_conjL
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1946
    fix W x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
    assume "openin X W" and "x \<in> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
    show "\<exists>U V. openin X U \<and> closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1949
    proof (cases "x=a")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
      have "compactin ?Xa (topspace X - W)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
        using \<open>openin X W\<close> assms(1) closedin_compact_space
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1953
        by (metis Diff_mono True \<open>x \<in> W\<close> compactin_subtopology empty_subsetI insert_subset openin_closedin_eq order_refl)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1954
      then obtain V K where V: "openin ?Xa V" and K: "compactin ?Xa K" "closedin ?Xa K" and "topspace X - W \<subseteq> V" "V \<subseteq> K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1955
        by (metis locally_compact_space_compact_closed_compact R)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
      proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
        show "openin X (topspace X - K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
          using "\<section>" K by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
        show "closedin X (topspace X - V)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
          using V ope openin_trans_full by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
        show "x \<in> topspace X - K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
        proof (rule)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1964
          show "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
            using \<open>openin X W\<close> \<open>x \<in> W\<close> openin_subset by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
          show "x \<notin> K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
            using K True closedin_imp_subset by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
        show "topspace X - K \<subseteq> topspace X - V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
          by (simp add: Diff_mono \<open>V \<subseteq> K\<close>)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1971
        show "topspace X - V \<subseteq> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
          using \<open>topspace X - W \<subseteq> V\<close> by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1973
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1974
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
      have "openin ?Xa ((topspace X - {a}) \<inter> W)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
        using \<open>openin X W\<close> openin_subtopology_Int2 by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
      moreover have "x \<in> (topspace X - {a}) \<inter> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
        using \<open>openin X W\<close> \<open>x \<in> W\<close> openin_subset False by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
      ultimately obtain U V where "openin ?Xa U" "compactin ?Xa V" "closedin ?Xa V"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
               "x \<in> U" "U \<subseteq> V" "V \<subseteq> (topspace X - {a}) \<inter> W"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
        using R locally_compact_regular_space_neighbourhood_base neighbourhood_base_of
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
        by (metis (no_types, lifting))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
        by (meson "\<section>" le_infE ope openin_trans_full)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1989
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
lemma regular_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
  "regular_space(Alexandroff_compactification X) \<longleftrightarrow> regular_space X \<and> locally_compact_space X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
  (is "regular_space ?Y = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
  have "regular_space ?Y \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
        regular_space (subtopology ?Y (topspace ?Y - {None})) \<and> locally_compact_space (subtopology ?Y (topspace ?Y - {None}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
    by (rule regular_space_one_point_compactification) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
  also have "... \<longleftrightarrow> regular_space X \<and> locally_compact_space X"
78200
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78128
diff changeset
  1999
    by (metis embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_locally_compact_space 
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78128
diff changeset
  2000
        homeomorphic_regular_space topspace_Alexandroff_compactification_delete)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
lemma Hausdorff_space_one_point_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
  assumes "compact_space X" and  "openin X (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
    and "\<And>K. \<lbrakk>compactin (subtopology X (topspace X - {a})) K; closedin (subtopology X (topspace X - {a})) K\<rbrakk> \<Longrightarrow> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
  shows "Hausdorff_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
           Hausdorff_space (subtopology X (topspace X - {a})) \<and> locally_compact_space (subtopology X (topspace X - {a}))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
    (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
  show ?rhs if ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
    have "locally_compact_space (subtopology X (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
      using assms that compact_imp_locally_compact_space locally_compact_space_open_subset 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
      by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
    with that show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
      by (simp add: Hausdorff_space_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
    by (metis assms locally_compact_Hausdorff_or_regular regular_space_one_point_compactification
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
        regular_t1_eq_Hausdorff_space t1_space_one_point_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2025
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
lemma Hausdorff_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
   "Hausdorff_space(Alexandroff_compactification X) \<longleftrightarrow> Hausdorff_space X \<and> locally_compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
  by (meson compact_Hausdorff_imp_regular_space compact_space_Alexandroff_compactification 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
      locally_compact_Hausdorff_or_regular regular_space_Alexandroff_compactification 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
      regular_t1_eq_Hausdorff_space t1_space_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2031
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
lemma completely_regular_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
   "completely_regular_space(Alexandroff_compactification X) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2034
        completely_regular_space X \<and> locally_compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
  by (metis regular_space_Alexandroff_compactification completely_regular_eq_regular_space
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
      compact_imp_locally_compact_space compact_space_Alexandroff_compactification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  2038
proposition Hausdorff_space_one_point_compactification_asymmetric_prod:
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
  shows "Hausdorff_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
         kc_space (prod_topology X (subtopology X (topspace X - {a}))) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
         k_space (prod_topology X (subtopology X (topspace X - {a})))"  (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
proof (cases "a \<in> topspace X")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
  proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
    show ?rhs if ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
    proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
      show "kc_space (prod_topology X (subtopology X (topspace X - {a})))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
        using Hausdorff_imp_kc_space kc_space_prod_topology_right kc_space_subtopology that by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2051
      show "k_space (prod_topology X (subtopology X (topspace X - {a})))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
        by (meson Hausdorff_imp_kc_space assms compact_imp_locally_compact_space k_space_prod_topology_left 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
            kc_space_one_point_compactification_gen that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
    assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
    define D where "D \<equiv> (\<lambda>x. (x,x)) ` (topspace X - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
    show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2059
    proof (cases "topspace X = {a}")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
        by (simp add: Hausdorff_space_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2064
      case False
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  2065
      with R True have "kc_space X"
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
        using kc_space_retraction_map_image [of "prod_topology X (subtopology X (topspace X - {a}))" X fst]
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  2067
        by (metis Diff_subset insert_Diff retraction_map_fst topspace_discrete_topology topspace_subtopology_subset)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2068
      have "closedin (subtopology (prod_topology X (subtopology X (topspace X - {a}))) K) (K \<inter> D)" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
        if "compactin (prod_topology X (subtopology X (topspace X - {a}))) K" for K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
      proof (intro closedin_subtopology_Int_subset[where V=K] closedin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
        show "fst ` K \<times> snd ` K \<inter> D \<subseteq> fst ` K \<times> snd ` K" "K \<subseteq> fst ` K \<times> snd ` K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2072
          by force+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2073
        have eq: "(fst ` K \<times> snd ` K \<inter> D) = ((\<lambda>x. (x,x)) ` (fst ` K \<inter> snd ` K))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
          using compactin_subset_topspace that by (force simp: D_def image_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
        have "compactin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K \<times> snd ` K \<inter> D)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
          unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
        proof (rule image_compactin [of "subtopology X (topspace X - {a})"])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
          have "compactin X (fst ` K)" "compactin X (snd ` K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
            by (meson compactin_subtopology continuous_map_fst continuous_map_snd image_compactin that)+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
          moreover have "fst ` K \<inter> snd ` K \<subseteq> topspace X - {a}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
            using compactin_subset_topspace that by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
          ultimately
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
          show "compactin (subtopology X (topspace X - {a})) (fst ` K \<inter> snd ` K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
            unfolding compactin_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
            by (meson \<open>kc_space X\<close> closed_Int_compactin kc_space_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
          show "continuous_map (subtopology X (topspace X - {a})) (prod_topology X (subtopology X (topspace X - {a}))) (\<lambda>x. (x,x))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2087
            by (simp add: continuous_map_paired)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2088
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
        then show "closedin (prod_topology X (subtopology X (topspace X - {a}))) (fst ` K \<times> snd ` K \<inter> D)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
          using R compactin_imp_closedin_gen by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
      moreover have "D \<subseteq> topspace X \<times> (topspace X \<inter> (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
        by (auto simp: D_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
      ultimately have *: "closedin (prod_topology X (subtopology X (topspace X - {a}))) D"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
        using R by (auto simp: k_space)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
      have "x=y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
        if "x \<in> topspace X" "y \<in> topspace X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
          and \<section>: "\<And>T. \<lbrakk>(x,y) \<in> T; openin (prod_topology X X) T\<rbrakk> \<Longrightarrow> \<exists>z \<in> topspace X. (z,z) \<in> T" for x y
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
      proof (cases "x=a \<and> y=a")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
        case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
        then consider "x\<noteq>a" | "y\<noteq>a"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
          by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
        then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
        proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
          case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
          have "\<exists>z \<in> topspace X - {a}. (z,z) \<in> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
            if "(y,x) \<in> T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
            have "(x,y) \<in> {z \<in> topspace (prod_topology X X). (snd z,fst z) \<in> T \<inter> topspace X \<times> (topspace X - {a})}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
              by (simp add: "1" \<open>x \<in> topspace X\<close> \<open>y \<in> topspace X\<close> that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2111
            moreover have "openin (prod_topology X X) {z \<in> topspace (prod_topology X X). (snd z,fst z) \<in> T \<inter> topspace X \<times> (topspace X - {a})}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
            proof (rule openin_continuous_map_preimage)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
              show "continuous_map (prod_topology X X) (prod_topology X X) (\<lambda>x. (snd x, fst x))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
                by (simp add: continuous_map_fst continuous_map_pairedI continuous_map_snd)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
              have "openin (prod_topology X X) (topspace X \<times> (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
                using \<open>kc_space X\<close> assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
              moreover have "openin (prod_topology X X) T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
                using kc_space_one_point_compactification_gen [OF \<open>compact_space X\<close>] \<open>kc_space X\<close> that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
                by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
              ultimately show "openin (prod_topology X X) (T \<inter> topspace X \<times> (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
                by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
            qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
            ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
              by (smt (verit) \<section> Int_iff fst_conv mem_Collect_eq mem_Sigma_iff snd_conv)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
          qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
          then have "(y,x) \<in> prod_topology X (subtopology X (topspace X - {a})) closure_of D"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
            by (simp add: "1" D_def in_closure_of that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
            using that "*" D_def closure_of_closedin by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
        next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
          case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
          have "\<exists>z \<in> topspace X - {a}. (z,z) \<in> T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
            if "(x,y) \<in> T" "openin (prod_topology X (subtopology X (topspace X - {a}))) T" for T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
            have "openin (prod_topology X X) (topspace X \<times> (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
              using \<open>kc_space X\<close> assms kc_space_one_point_compactification_gen openin_prod_Times_iff by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2137
            moreover have XXT: "openin (prod_topology X X) T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
              using kc_space_one_point_compactification_gen [OF \<open>compact_space X\<close>] \<open>kc_space X\<close> that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2139
              by (metis openin_prod_Times_iff openin_topspace openin_trans_full prod_topology_subtopology(2))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2140
            ultimately have "openin (prod_topology X X) (T \<inter> topspace X \<times> (topspace X - {a}))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
              by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
            then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
              by (smt (verit) "\<section>" Diff_subset XXT mem_Sigma_iff openin_subset subsetD that topspace_prod_topology topspace_subtopology_subset)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
          qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
          then have "(x,y) \<in> prod_topology X (subtopology X (topspace X - {a})) closure_of D"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
            by (simp add: "2" D_def in_closure_of that)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
            using that "*" D_def closure_of_closedin by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2150
      qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
        unfolding Hausdorff_space_closedin_diagonal closure_of_subset_eq [symmetric] 
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2153
          by (force simp: closure_of_def)
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
  case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
    by (simp add: assms compact_imp_k_space compact_space_prod_topology kc_space_compact_prod_topology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2162
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
lemma Hausdorff_space_Alexandroff_compactification_asymmetric_prod:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2164
   "Hausdorff_space(Alexandroff_compactification X) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
        kc_space(prod_topology (Alexandroff_compactification X) X) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
        k_space(prod_topology (Alexandroff_compactification X) X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
    (is "Hausdorff_space ?Y = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
  have *: "subtopology (Alexandroff_compactification X)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
     (topspace (Alexandroff_compactification X) -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
      {None}) homeomorphic_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
    using embedding_map_Some embedding_map_imp_homeomorphic_space homeomorphic_space_sym by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
  have "Hausdorff_space (Alexandroff_compactification X) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2174
      (kc_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
       k_space (prod_topology ?Y (subtopology ?Y (topspace ?Y - {None}))))"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
    by (rule Hausdorff_space_one_point_compactification_asymmetric_prod) (auto simp: compactin_subtopology closedin_subtopology closedin_Alexandroff_I)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
  also have "... \<longleftrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
    using homeomorphic_k_space homeomorphic_kc_space homeomorphic_space_prod_topology 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
          homeomorphic_space_refl * by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2182
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2183
lemma kc_space_as_compactification_unique:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
  assumes "kc_space X" "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2185
  shows "openin X U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
         (if a \<in> U then U \<subseteq> topspace X \<and> compactin X (topspace X - U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
                   else openin (subtopology X (topspace X - {a})) U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2188
proof (cases "a \<in> U")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2189
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
    by (meson assms closedin_compact_space compactin_imp_closedin_gen openin_closedin_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
  case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
  by (metis Diff_empty kc_space_one_point_compactification_gen openin_open_subtopology openin_subset subset_Diff_insert assms)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2197
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
lemma kc_space_as_compactification_unique_explicit:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
  assumes "kc_space X" "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
  shows "openin X U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
         (if a \<in> U then U \<subseteq> topspace X \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
                     compactin (subtopology X (topspace X - {a})) (topspace X - U) \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
                     closedin (subtopology X (topspace X - {a})) (topspace X - U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
                else openin (subtopology X (topspace X - {a})) U)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2205
  apply (simp add: kc_space_subtopology compactin_imp_closedin_gen assms compactin_subtopology cong: conj_cong)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
  by (metis Diff_mono assms bot.extremum insert_subset kc_space_as_compactification_unique subset_refl)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
lemma Alexandroff_compactification_unique:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
  assumes "kc_space X" "compact_space X" and a: "a \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
  shows "Alexandroff_compactification (subtopology X (topspace X - {a})) homeomorphic_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
        (is "?Y homeomorphic_space X")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
  have [simp]: "topspace X \<inter> (topspace X - {a}) = topspace X - {a}"  
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
  have [simp]: "insert None (Some ` A) = insert None (Some ` B) \<longleftrightarrow> A = B" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
               "insert None (Some ` A) \<noteq> Some ` B" for A B
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
  have "quotient_map X ?Y (\<lambda>x. if x = a then None else Some x)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
    unfolding quotient_map_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
    show "(\<lambda>x. if x = a then None else Some x) ` topspace X = topspace ?Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
      using \<open>a \<in> topspace X\<close>  by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
    show "openin X {x \<in> topspace X. (if x = a then None else Some x) \<in> U} = openin ?Y U" (is "?L = ?R")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2224
      if "U \<subseteq> topspace ?Y" for U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2225
    proof (cases "None \<in> U")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
      then obtain T where T[simp]: "U = insert None (Some ` T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
        by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
      have Tsub: "T \<subseteq> topspace X - {a}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
        using in_these_eq that by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2231
      then have "{x \<in> topspace X. (if x = a then None else Some x) \<in> U} = insert a T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
        by (auto simp: a image_iff cong: conj_cong)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
      then have "?L \<longleftrightarrow> openin X (insert a T)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
        by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
      also have "\<dots> \<longleftrightarrow> ?R"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
        using Tsub assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
        apply (simp add: openin_Alexandroff_compactification kc_space_as_compactification_unique_explicit [where a=a])
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
        by (smt (verit, best) Diff_insert2 Diff_subset closedin_imp_subset double_diff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
      finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
      then obtain T where [simp]: "U = Some ` T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
        by (metis Int_insert_right UNIV_I UNIV_option_conv inf.orderE inf_le2 subsetI subset_imageE)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
      have **: "\<And>V. openin X V \<Longrightarrow> openin X (V - {a})"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2245
        by (simp add: assms compactin_imp_closedin_gen openin_diff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
      have Tsub: "T \<subseteq> topspace X - {a}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
        using in_these_eq that by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
      then have "{x \<in> topspace X. (if x = a then None else Some x) \<in> U} = T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
        by (auto simp: image_iff cong: conj_cong)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
        by (simp add: "**" Tsub openin_open_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2253
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
  moreover have "inj_on (\<lambda>x. if x = a then None else Some x) (topspace X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
    by (auto simp: inj_on_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2256
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
    using homeomorphic_space_sym homeomorphic_space homeomorphic_map_def by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2260
subsection\<open>Extending continuous maps "pointwise" in a regular space\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2261
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2262
lemma continuous_map_on_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2263
  assumes Y: "regular_space Y" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2264
    and T: "T \<subseteq> X closure_of S" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2265
    and f: "\<And>t. t \<in> T \<Longrightarrow> limitin Y f (f t) (atin_within X t S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2266
  shows "continuous_map (subtopology X T) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2267
proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2268
  fix a
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2269
  assume "a \<in> topspace X" and "a \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2270
  have "f ` T \<subseteq> topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2271
    by (metis f image_subsetI limitin_topspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2272
  have "\<forall>\<^sub>F x in atin_within X a T. f x \<in> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2273
    if W: "openin Y W" "f a \<in> W" for W
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2274
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2275
    obtain V C where "openin Y V" "closedin Y C" "f a \<in> V" "V \<subseteq> C" "C \<subseteq> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2276
      by (metis Y W neighbourhood_base_of neighbourhood_base_of_closedin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2277
    have "\<forall>\<^sub>F x in atin_within X a S. f x \<in> V"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2278
      by (metis \<open>a \<in> T\<close> \<open>f a \<in> V\<close> \<open>openin Y V\<close> f limitin_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2279
    then obtain U where "openin X U" "a \<in> U" and U: "\<forall>x \<in> U - {a}. x \<in> S \<longrightarrow> f x \<in> V"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2280
      by (smt (verit) Diff_iff \<open>a \<in> topspace X\<close> eventually_atin_within insert_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2281
    moreover have "f z \<in> W" if "z \<in> U" "z \<noteq> a" "z \<in> T" for z
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2282
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2283
      have "z \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2284
        using \<open>openin X U\<close> openin_subset \<open>z \<in> U\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2285
      then have "f z \<in> topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2286
        using \<open>f ` T \<subseteq> topspace Y\<close> \<open>z \<in> T\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2287
      { assume "f z \<in> topspace Y" "f z \<notin> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2288
        then have "\<forall>\<^sub>F x in atin_within X z S. f x \<in> topspace Y - C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2289
          by (metis Diff_iff \<open>closedin Y C\<close> closedin_def f limitinD \<open>z \<in> T\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2290
        then obtain U' where U': "openin X U'" "z \<in> U'" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2291
                 "\<And>x. x \<in> U' - {z} \<Longrightarrow> x \<in> S \<Longrightarrow> f x \<notin> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2292
          by (smt (verit) Diff_iff \<open>z \<in> topspace X\<close> eventually_atin_within insertCI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2293
        then have *: "\<And>D. z \<in> D \<and> openin X D \<Longrightarrow> \<exists>y. y \<in> S \<and> y \<in> D"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2294
          by (meson T in_closure_of subsetD \<open>z \<in> T\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2295
        have False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2296
          using * [of "U \<inter> U'"] U' U \<open>V \<subseteq> C\<close> \<open>f a \<in> V\<close> \<open>f z \<notin> C\<close> \<open>openin X U\<close> that
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2297
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2298
      }
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2299
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2300
        using \<open>C \<subseteq> W\<close> \<open>f z \<in> topspace Y\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2301
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2302
    ultimately have "\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. x \<in> T \<longrightarrow> f x \<in> W)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2303
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2304
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2305
      using eventually_atin_within by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2306
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2307
  then show "limitin Y f (f a) (atin (subtopology X T) a)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2308
    by (metis \<open>a \<in> T\<close> atin_subtopology_within f limitin_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2309
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2310
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2311
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2312
lemma continuous_map_on_intermediate_closure_of_eq:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2313
  assumes "regular_space Y" "S \<subseteq> T" and Tsub: "T \<subseteq> X closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2314
  shows "continuous_map (subtopology X T) Y f \<longleftrightarrow> (\<forall>t \<in> T. limitin Y f (f t) (atin_within X t S))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2315
        (is "?lhs \<longleftrightarrow> ?rhs")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2316
proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2317
  assume L: ?lhs
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2318
  show ?rhs
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2319
  proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2320
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2321
    assume "x \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2322
    with L Tsub closure_of_subset_topspace 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2323
    have "limitin Y f (f x) (atin (subtopology X T) x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2324
      by (fastforce simp: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2325
    then show "limitin Y f (f x) (atin_within X x S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2326
      using \<open>x \<in> T\<close> \<open>S \<subseteq> T\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2327
      by (force simp: limitin_def atin_subtopology_within eventually_atin_within)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2328
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2329
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2330
  show "?rhs \<Longrightarrow> ?lhs" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2331
    using assms continuous_map_on_intermediate_closure_of by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2332
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2333
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2334
lemma continuous_map_extension_pointwise_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2335
  assumes \<section>: "regular_space Y" "S \<subseteq> T" "T \<subseteq> X closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2336
    and f: "continuous_map (subtopology X S) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2337
    and lim: "\<And>t. t \<in> T-S \<Longrightarrow> \<exists>l. limitin Y f l (atin_within X t S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2338
  obtains g where "continuous_map (subtopology X T) Y g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2339
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2340
  obtain g where g: "\<And>t. t \<in> T \<and> t \<notin> S \<Longrightarrow> limitin Y f (g t) (atin_within X t S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2341
    by (metis Diff_iff lim)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2342
  let ?h = "\<lambda>x. if x \<in> S then f x else g x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2343
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2344
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2345
    have T: "T \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2346
      using \<section> closure_of_subset_topspace by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2347
    have "limitin Y ?h (f t) (atin_within X t S)" if "t \<in> T" "t \<in> S" for t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2348
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2349
      have "limitin Y f (f t) (atin_within X t S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2350
        by (meson T f limit_continuous_map_within subset_eq that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2351
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2352
        by (simp add: eventually_atin_within limitin_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2353
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2354
    moreover have "limitin Y ?h (g t) (atin_within X t S)" if "t \<in> T" "t \<notin> S" for t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2355
      by (smt (verit, del_insts) eventually_atin_within g limitin_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2356
    ultimately show "continuous_map (subtopology X T) Y ?h"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2357
      unfolding continuous_map_on_intermediate_closure_of_eq [OF \<section>] 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2358
      by (auto simp: \<section> atin_subtopology_within)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2359
  qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2360
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2361
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2362
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2363
lemma continuous_map_extension_pointwise:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2364
  assumes "regular_space Y" "S \<subseteq> T" and Tsub: "T \<subseteq> X closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2365
    and ex: " \<And>x. x \<in> T \<Longrightarrow> \<exists>g. continuous_map (subtopology X (insert x S)) Y g \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2366
                     (\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2367
  obtains g where "continuous_map (subtopology X T) Y g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2368
proof (rule continuous_map_extension_pointwise_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2369
  show "continuous_map (subtopology X S) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2370
  proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2371
    fix t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2372
    assume "t \<in> topspace X" and "t \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2373
    then obtain g where g: "limitin Y g (g t) (atin (subtopology X (insert t S)) t)" and gf: "\<forall>x \<in> S. g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2374
      by (metis Int_iff \<open>S \<subseteq> T\<close> continuous_map_atin ex inf.orderE insert_absorb topspace_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2375
    with \<open>t \<in> S\<close> show "limitin Y f (f t) (atin (subtopology X S) t)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2376
      by (simp add: limitin_def atin_subtopology_within_if eventually_atin_within gf insert_absorb)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2377
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2378
  show "\<exists>l. limitin Y f l (atin_within X t S)" if "t \<in> T - S" for t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2379
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2380
    obtain g where g: "continuous_map (subtopology X (insert t S)) Y g" and gf: "\<forall>x \<in> S. g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2381
      using \<open>S \<subseteq> T\<close> ex \<open>t \<in> T - S\<close> by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2382
    then have "limitin Y g (g t) (atin_within X t (insert t S))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2383
      using Tsub in_closure_of limit_continuous_map_within that  by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2384
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2385
      unfolding limitin_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2386
      by (smt (verit) eventually_atin_within gf subsetD subset_insertI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2387
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2388
qed (use assms in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2389
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2390
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2391
subsection\<open>Extending Cauchy continuous functions to the closure\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2392
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2393
lemma Cauchy_continuous_map_extends_to_continuous_closure_of_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2394
  assumes m2: "mcomplete_of m2" and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2395
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2396
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2397
  where "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2398
                        (mtopology_of m2) g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2399
proof (rule continuous_map_extension_pointwise_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2400
  interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2401
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2402
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2403
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2404
  show "regular_space (mtopology_of m2)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2405
    by (simp add: Metric_space.regular_space_mtopology mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2406
  show "S \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2407
    by (simp add: assms(3) closure_of_subset)    
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2408
  show "continuous_map (subtopology (mtopology_of m1) S) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2409
    by (metis Cauchy_continuous_imp_continuous_map f mtopology_of_submetric)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2410
  fix a
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2411
  assume a: "a \<in> mtopology_of m1 closure_of S - S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2412
  then obtain \<sigma> where ran\<sigma>: "range \<sigma> \<subseteq> S" "range \<sigma> \<subseteq> mspace m1" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2413
    and lim\<sigma>: "limitin L.M1.mtopology \<sigma> a sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2414
    by (force simp: mtopology_of_def L.M1.closure_of_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2415
  then have "L.M1.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2416
    by (simp add: L.M1.convergent_imp_MCauchy mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2417
  then have "L.M2.MCauchy (f \<circ> \<sigma>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2418
    using f ran\<sigma> by (simp add: Cauchy_continuous_map_def L.M1.subspace Metric_space.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2419
  then obtain l where l: "limitin L.M2.mtopology (f \<circ> \<sigma>) l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2420
    by (meson L.M2.mcomplete_def m2 mcomplete_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2421
  have "limitin L.M2.mtopology f l (atin_within L.M1.mtopology a S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2422
    unfolding L.limit_atin_sequentially_within imp_conjL
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2423
  proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2424
    show "l \<in> mspace m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2425
      using L.M2.limitin_mspace l by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2426
    fix \<rho>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2427
    assume "range \<rho> \<subseteq> S \<inter> mspace m1 - {a}" and lim\<rho>: "limitin L.M1.mtopology \<rho> a sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2428
    then have ran\<rho>: "range \<rho> \<subseteq> S" "range \<rho> \<subseteq> mspace m1" "\<And>n. \<rho> n \<noteq> a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2429
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2430
    have "a \<in> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2431
      using L.M1.limitin_mspace lim\<rho> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2432
    have "S.MCauchy \<sigma>" "S.MCauchy \<rho>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2433
      using L.M1.convergent_imp_MCauchy L.M1.MCauchy_def S.MCauchy_def lim\<sigma> ran\<sigma> lim\<rho> ran\<rho> by force+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2434
    then have "L.M2.MCauchy (f \<circ> \<rho>)" "L.M2.MCauchy (f \<circ> \<sigma>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2435
      using f by (auto simp: Cauchy_continuous_map_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2436
    then have ran_f: "range (\<lambda>x. f (\<rho> x)) \<subseteq> mspace m2" "range (\<lambda>x. f (\<sigma> x)) \<subseteq> mspace m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2437
      by (auto simp: L.M2.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2438
    have "(\<lambda>n. mdist m2 (f (\<rho> n)) l) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2439
    proof (rule Lim_null_comparison)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2440
      have "mdist m2 (f (\<rho> n)) l \<le> mdist m2 (f (\<sigma> n)) l + mdist m2 (f (\<sigma> n)) (f (\<rho> n))" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2441
        using \<open>l \<in> mspace m2\<close> ran_f L.M2.triangle'' by (smt (verit, best) range_subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2442
      then show "\<forall>\<^sub>F n in sequentially. norm (mdist m2 (f (\<rho> n)) l) \<le> mdist m2 (f (\<sigma> n)) l + mdist m2 (f (\<sigma> n)) (f (\<rho> n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2443
        by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2444
      define \<psi> where "\<psi> \<equiv> \<lambda>n. if even n then \<sigma> (n div 2) else \<rho> (n div 2)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2445
      have "(\<lambda>n. mdist m1 (\<sigma> n) (\<rho> n)) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2446
      proof (rule Lim_null_comparison)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2447
        show "\<forall>\<^sub>F n in sequentially. norm (mdist m1 (\<sigma> n) (\<rho> n)) \<le> mdist m1 (\<sigma> n) a + mdist m1 (\<rho> n) a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2448
          using L.M1.triangle' [of _ a] ran\<sigma> ran\<rho> \<open>a \<in> mspace m1\<close> by (simp add: range_subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2449
        have "(\<lambda>n. mdist m1 (\<sigma> n) a) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2450
          using L.M1.limitin_metric_dist_null lim\<sigma> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2451
        moreover have "(\<lambda>n. mdist m1 (\<rho> n) a) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2452
          using L.M1.limitin_metric_dist_null lim\<rho> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2453
        ultimately show "(\<lambda>n. mdist m1 (\<sigma> n) a + mdist m1 (\<rho> n) a) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2454
          by (simp add: tendsto_add_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2455
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2456
      with \<open>S.MCauchy \<sigma>\<close> \<open>S.MCauchy \<rho>\<close> have "S.MCauchy \<psi>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2457
        by (simp add: S.MCauchy_interleaving_gen \<psi>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2458
      then have "L.M2.MCauchy (f \<circ> \<psi>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2459
        by (metis Cauchy_continuous_map_def f mdist_submetric mspace_submetric)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2460
      then have "(\<lambda>n. mdist m2 (f (\<sigma> n)) (f (\<rho> n))) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2461
        using L.M2.MCauchy_interleaving_gen [of "f \<circ> \<sigma>" "f \<circ> \<rho>"]  
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2462
        by (simp add: if_distrib \<psi>_def o_def cong: if_cong)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2463
      moreover have "\<forall>\<^sub>F n in sequentially. f (\<sigma> n) \<in> mspace m2 \<and> (\<lambda>x. mdist m2 (f (\<sigma> x)) l) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2464
        using l by (auto simp: L.M2.limitin_metric_dist_null \<open>l \<in> mspace m2\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2465
      ultimately show "(\<lambda>n. mdist m2 (f (\<sigma> n)) l + mdist m2 (f (\<sigma> n)) (f (\<rho> n))) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2466
        by (metis (mono_tags) tendsto_add_zero eventually_sequentially order_refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2467
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2468
    with ran_f show "limitin L.M2.mtopology (f \<circ> \<rho>) l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2469
      by (auto simp: L.M2.limitin_metric_dist_null eventually_sequentially \<open>l \<in> mspace m2\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2470
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2471
  then show "\<exists>l. limitin (mtopology_of m2) f l (atin_within (mtopology_of m1) a S)" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2472
    by (force simp: mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2473
qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2474
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2475
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2476
lemma Cauchy_continuous_map_extends_to_continuous_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2477
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2478
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2479
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2480
  where "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2481
                        (mtopology_of m2) g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2482
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2483
  obtain g where cmg: 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2484
    "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of (mspace m1 \<inter> S))) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2485
                        (mtopology_of m2) g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2486
    and gf: "(\<forall>x \<in> mspace m1 \<inter> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2487
    using Cauchy_continuous_map_extends_to_continuous_closure_of_aux assms
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2488
    by (metis inf_commute inf_le2 submetric_restrict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2489
  define h where "h \<equiv> \<lambda>x. if x \<in> topspace(mtopology_of m1) then g x else f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2490
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2491
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2492
    show "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2493
                         (mtopology_of m2) h"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2494
      unfolding h_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2495
    proof (rule continuous_map_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2496
      show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2497
        by (metis closure_of_restrict cmg topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2498
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2499
  qed (auto simp: gf h_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2500
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2501
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2502
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2503
lemma Cauchy_continuous_map_extends_to_continuous_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2504
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2505
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2506
    and T: "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2507
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2508
  where "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2509
         "(\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2510
  by (metis Cauchy_continuous_map_extends_to_continuous_closure_of T assms(1) continuous_map_from_subtopology_mono f)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2511
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2512
text \<open>Technical lemma helpful for porting particularly ugly HOL Light proofs\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2513
lemma all_in_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2514
  assumes P: "\<forall>x \<in> S. P x" and clo: "closedin X {x \<in> topspace X. P x}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2515
  shows "\<forall>x \<in> X closure_of S. P x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2516
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2517
  have *: "topspace X \<inter> S \<subseteq> {x \<in> topspace X. P x}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2518
    using P by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2519
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2520
  using closure_of_minimal [OF * clo]  closure_of_restrict by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2521
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2522
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2523
lemma Lipschitz_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2524
  assumes lcf: "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2525
    and "S \<subseteq> T" and Tsub: "T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2526
    and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2527
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2528
  shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2529
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2530
  interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2531
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2532
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2533
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2534
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2535
    using Tsub by (auto simp: mtopology_of_def closure_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2536
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2537
    unfolding Lipschitz_continuous_map_pos
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2538
  proof
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2539
    show "f \<in> mspace (submetric m1 T) \<rightarrow> mspace m2"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2540
      by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist  mtopology_of_def 
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2541
          mtopology_of_submetric image_subset_iff_funcset)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2542
    define X where "X \<equiv> prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2543
    obtain B::real where "B > 0" and B: "\<forall>(x,y) \<in> S\<times>S. mdist m2 (f x) (f y) \<le> B * mdist m1 x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2544
      using lcf \<open>S \<subseteq> mspace m1\<close>  by (force simp: Lipschitz_continuous_map_pos)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2545
    have eq: "{z \<in> A. case z of (x,y) \<Rightarrow> p x y \<le> B * q x y} = {z \<in> A. ((\<lambda>(x,y). B * q x y - p x y)z) \<in> {0..}}" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2546
        for p q and A::"('a*'a)set"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2547
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2548
    have clo: "closedin X {z \<in> topspace X. case z of (x, y) \<Rightarrow> mdist m2 (f x) (f y) \<le> B * mdist m1 x y}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2549
      unfolding eq
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2550
    proof (rule closedin_continuous_map_preimage)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2551
      have *: "continuous_map X L.M2.mtopology (f \<circ> fst)" "continuous_map X L.M2.mtopology (f \<circ> snd)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2552
        using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2553
      then show "continuous_map X euclidean (\<lambda>x. case x of (x, y) \<Rightarrow> B * mdist m1 x y - mdist m2 (f x) (f y))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2554
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2555
      proof (intro continuous_intros; simp add: mtopology_of_def o_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2556
        show "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2557
          by (simp_all add: X_def continuous_map_subtopology_fst continuous_map_subtopology_snd flip: subtopology_Times)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2558
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2559
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2560
    have "mdist m2 (f x) (f y) \<le> B * mdist m1 x y" if "x \<in> T" "y \<in> T" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2561
      using all_in_closure_of [OF B clo] \<open>S \<subseteq> T\<close> Tsub
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2562
      by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2  
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2563
          mtopology_of_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2564
    then show "\<exists>B>0. \<forall>x\<in>mspace (submetric m1 T).
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2565
             \<forall>y\<in>mspace (submetric m1 T).
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2566
                mdist m2 (f x) (f y) \<le> B * mdist (submetric m1 T) x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2567
      using \<open>0 < B\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2568
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2569
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2570
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2571
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2572
lemma Lipschitz_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2573
  assumes "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2574
    and "S \<subseteq> T" "T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2575
    and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2576
  shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2577
  by (metis Lipschitz_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace subset_trans topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2578
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2579
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2580
lemma Lipschitz_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2581
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2582
    and f: "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2583
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2584
  where "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2585
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2586
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2587
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2588
    where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2589
                        (mtopology_of m2) g"  "(\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2590
    by (metis Cauchy_continuous_map_extends_to_continuous_closure_of Lipschitz_imp_Cauchy_continuous_map f m2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2591
  have "Lipschitz_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2592
  proof (rule Lipschitz_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2593
    show "Lipschitz_continuous_map (submetric m1 (mspace m1 \<inter> S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2594
      by (smt (verit, best) IntD2 Lipschitz_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2595
    show "mspace m1 \<inter> S \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2596
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2597
    show "mtopology_of m1 closure_of S \<subseteq> mtopology_of m1 closure_of (mspace m1 \<inter> S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2598
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2599
    show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2600
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2601
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2602
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2603
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2604
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2605
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2606
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2607
lemma Lipschitz_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2608
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2609
    and "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2610
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2611
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2612
  where "Lipschitz_continuous_map (submetric m1 T) m2 g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2613
  by (metis Lipschitz_continuous_map_extends_to_closure_of Lipschitz_continuous_map_from_submetric_mono assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2614
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2615
text \<open>This proof uses the same trick to extend the function's domain to its closure\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2616
lemma uniformly_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2617
  assumes ucf: "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2618
    and "S \<subseteq> T" and Tsub: "T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2619
    and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2620
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2621
  shows "uniformly_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2622
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2623
  interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2624
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2625
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2626
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2627
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2628
    using Tsub by (auto simp: mtopology_of_def closure_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2629
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2630
    unfolding uniformly_continuous_map_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2631
    proof (intro conjI strip)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2632
    show "f \<in> mspace (submetric m1 T) \<rightarrow> mspace m2"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2633
      by (metis cmf Metric_space.metric_continuous_map Metric_space_mspace_mdist  
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2634
          mtopology_of_def mtopology_of_submetric image_subset_iff_funcset)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2635
    fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2636
    assume "\<epsilon> > 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2637
    then obtain \<delta> where "\<delta>>0" and \<delta>: "\<forall>(x,y) \<in> S\<times>S. mdist m1 x y < \<delta> \<longrightarrow> mdist m2 (f x) (f y) \<le> \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2638
      using ucf \<open>S \<subseteq> mspace m1\<close> unfolding uniformly_continuous_map_def mspace_submetric
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2639
      apply (simp add: Ball_def del: divide_const_simps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2640
      by (metis IntD2 half_gt_zero inf.orderE less_eq_real_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2641
    define X where "X \<equiv> prod_topology (subtopology L.M1.mtopology T) (subtopology L.M1.mtopology T)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2642
    text \<open>A clever construction involving the union of two closed sets\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2643
    have eq: "{z \<in> A. case z of (x,y) \<Rightarrow> p x y < d \<longrightarrow> q x y \<le> e} 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2644
            = {z \<in> A. ((\<lambda>(x,y). p x y - d)z) \<in> {0..}} \<union> {z \<in> A. ((\<lambda>(x,y). e - q x y)z) \<in> {0..}}" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2645
      for p q and d e::real and A::"('a*'a)set"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2646
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2647
    have clo: "closedin X {z \<in> topspace X. case z of (x, y) \<Rightarrow> mdist m1 x y < \<delta> \<longrightarrow> mdist m2 (f x) (f y) \<le> \<epsilon>/2}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2648
      unfolding eq
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2649
    proof (intro closedin_Un closedin_continuous_map_preimage)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2650
      have *: "continuous_map X L.M1.mtopology fst" "continuous_map X L.M1.mtopology snd"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2651
        by (metis X_def continuous_map_subtopology_fst subtopology_Times continuous_map_subtopology_snd)+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2652
      show "continuous_map X euclidean (\<lambda>x. case x of (x, y) \<Rightarrow> mdist m1 x y - \<delta>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2653
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2654
        by (intro continuous_intros; simp add: mtopology_of_def *)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2655
      have *: "continuous_map X L.M2.mtopology (f \<circ> fst)" "continuous_map X L.M2.mtopology (f \<circ> snd)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2656
        using cmf by (auto simp: mtopology_of_def X_def intro: continuous_map_compose continuous_map_fst continuous_map_snd)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2657
      then show "continuous_map X euclidean (\<lambda>x. case x of (x, y) \<Rightarrow> \<epsilon> / 2 - mdist m2 (f x) (f y))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2658
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2659
        by (intro continuous_intros; simp add: mtopology_of_def o_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2660
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2661
    have "mdist m2 (f x) (f y) \<le> \<epsilon>/2" if "x \<in> T" "y \<in> T" "mdist m1 x y < \<delta>" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2662
      using all_in_closure_of [OF \<delta> clo] \<open>S \<subseteq> T\<close> Tsub
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2663
      by (fastforce simp: X_def subset_iff closure_of_Times closure_of_subtopology inf.absorb2  
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2664
          mtopology_of_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2665
    then show "\<exists>\<delta>>0. \<forall>x\<in>mspace (submetric m1 T). \<forall>y\<in>mspace (submetric m1 T). mdist (submetric m1 T) y x < \<delta> \<longrightarrow> mdist m2 (f y) (f x) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2666
      using \<open>0 < \<delta>\<close> \<open>0 < \<epsilon>\<close> by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2667
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2668
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2669
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2670
lemma uniformly_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2671
  assumes "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2672
    and "S \<subseteq> T" and"T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2673
    and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2674
  shows "uniformly_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2675
  by (metis assms closure_of_subset_topspace subset_trans topspace_mtopology_of 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2676
      uniformly_continuous_map_on_intermediate_closure_aux)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2677
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2678
lemma uniformly_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2679
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2680
    and f: "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2681
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2682
  where "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2683
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2684
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2685
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2686
    where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2687
                        (mtopology_of m2) g"  "(\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2688
    by (metis Cauchy_continuous_map_extends_to_continuous_closure_of uniformly_imp_Cauchy_continuous_map f m2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2689
  have "uniformly_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2690
  proof (rule uniformly_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2691
    show "uniformly_continuous_map (submetric m1 (mspace m1 \<inter> S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2692
      by (smt (verit, best) IntD2 uniformly_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2693
    show "mspace m1 \<inter> S \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2694
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2695
    show "mtopology_of m1 closure_of S \<subseteq> mtopology_of m1 closure_of (mspace m1 \<inter> S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2696
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2697
    show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2698
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2699
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2700
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2701
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2702
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2703
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2704
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2705
lemma uniformly_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2706
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2707
    and "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2708
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2709
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2710
  where "uniformly_continuous_map (submetric m1 T) m2 g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2711
  by (metis uniformly_continuous_map_extends_to_closure_of uniformly_continuous_map_from_submetric_mono assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2712
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2713
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2714
lemma Cauchy_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2715
  assumes ucf: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2716
    and "S \<subseteq> T" and Tsub: "T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2717
    and cmf: "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2718
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2719
  shows "Cauchy_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2720
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2721
  interpret L: Metric_space12 "mspace m1" "mdist m1" "mspace m2" "mdist m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2722
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2723
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2724
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2725
  interpret T: Metric_space T "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2726
    by (metis L.M1.subspace Tsub closure_of_subset_topspace dual_order.trans topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2727
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2728
    using Tsub by (auto simp: mtopology_of_def closure_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2729
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2730
  proof (clarsimp simp: Cauchy_continuous_map_def Int_absorb2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2731
    fix \<sigma>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2732
    assume \<sigma>: "T.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2733
    have "\<exists>y\<in>S. mdist m1 (\<sigma> n) y < inverse (Suc n) \<and> mdist m2 (f (\<sigma> n)) (f y) < inverse (Suc n)" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2734
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2735
      have "\<sigma> n \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2736
        using \<sigma> by (force simp: T.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2737
      moreover have "continuous_map (mtopology_of (submetric m1 T)) L.M2.mtopology f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2738
        by (metis cmf mtopology_of_def mtopology_of_submetric)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2739
      ultimately obtain \<delta> where "\<delta>>0" and \<delta>: "\<forall>x \<in> T. mdist m1 (\<sigma> n) x < \<delta> \<longrightarrow> mdist m2 (f(\<sigma> n)) (f x) < inverse (Suc n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2740
        using \<open>T \<subseteq> mspace m1\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2741
        apply (simp add: mtopology_of_def Metric_space.metric_continuous_map L.M1.subspace Int_absorb2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2742
        by (metis inverse_Suc of_nat_Suc)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2743
      have "\<exists>y \<in> S. mdist m1 (\<sigma> n) y < min \<delta> (inverse (Suc n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2744
        using \<open>\<sigma> n \<in> T\<close> Tsub \<open>\<delta>>0\<close> 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2745
        unfolding mtopology_of_def L.M1.metric_closure_of subset_iff mem_Collect_eq L.M1.in_mball
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2746
        by (smt (verit, del_insts) inverse_Suc )
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2747
      with \<delta> \<open>S \<subseteq> T\<close> show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2748
        by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2749
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2750
    then obtain \<rho> where \<rho>S: "\<And>n. \<rho> n \<in> S" and \<rho>1: "\<And>n. mdist m1 (\<sigma> n) (\<rho> n) < inverse (Suc n)" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2751
                    and \<rho>2: "\<And>n. mdist m2 (f (\<sigma> n)) (f (\<rho> n)) < inverse (Suc n)" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2752
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2753
    have "S.MCauchy \<rho>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2754
      unfolding S.MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2755
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2756
      show "range \<rho> \<subseteq> S \<inter> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2757
        using \<open>S \<subseteq> mspace m1\<close> by (auto simp: \<rho>S)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2758
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2759
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2760
      then obtain M where M: "\<And>n n'. M \<le> n \<Longrightarrow> M \<le> n' \<Longrightarrow> mdist m1 (\<sigma> n) (\<sigma> n') < \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2761
        using \<sigma> unfolding T.MCauchy_def by (meson half_gt_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2762
      have "\<forall>\<^sub>F n in sequentially. inverse (Suc n) < \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2763
        using Archimedean_eventually_inverse \<open>0 < \<epsilon>\<close> divide_pos_pos zero_less_numeral by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2764
      then obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> inverse (Suc n) < \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2765
        by (meson eventually_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2766
      have "mdist m1 (\<rho> n) (\<rho> n') < \<epsilon>" if "n \<ge> max M N" "n' \<ge> max M N" for n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2767
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2768
        have "mdist m1 (\<rho> n) (\<rho> n') \<le> mdist m1 (\<rho> n) (\<sigma> n) + mdist m1 (\<sigma> n) (\<rho> n')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2769
          by (meson T.MCauchy_def T.triangle \<rho>S \<sigma> \<open>S \<subseteq> T\<close> rangeI subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2770
        also have "\<dots> \<le> mdist m1 (\<rho> n) (\<sigma> n) + mdist m1 (\<sigma> n) (\<sigma> n') + mdist m1 (\<sigma> n') (\<rho> n')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2771
          by (smt (verit, best) T.MCauchy_def T.triangle \<rho>S \<sigma> \<open>S \<subseteq> T\<close> in_mono rangeI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2772
        also have "\<dots> < \<epsilon>/4 + \<epsilon>/2 + \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2773
          using \<rho>1[of n] \<rho>1[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: T.commute)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2774
        also have "... \<le> \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2775
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2776
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2777
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2778
      then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> mdist m1 (\<rho> n) (\<rho> n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2779
        by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2780
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2781
    then have f\<rho>: "L.M2.MCauchy (f \<circ> \<rho>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2782
      using ucf by (simp add: Cauchy_continuous_map_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2783
    show "L.M2.MCauchy (f \<circ> \<sigma>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2784
      unfolding L.M2.MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2785
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2786
      show "range (f \<circ> \<sigma>) \<subseteq> mspace m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2787
        using \<open>T \<subseteq> mspace m1\<close> \<sigma> cmf
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2788
        apply (auto simp: )
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2789
        by (metis Metric_space.metric_continuous_map Metric_space_mspace_mdist T.MCauchy_def image_eqI inf.absorb1 mspace_submetric mtopology_of_def mtopology_of_submetric range_subsetD subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2790
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2791
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2792
      then obtain M where M: "\<And>n n'. M \<le> n \<Longrightarrow> M \<le> n' \<Longrightarrow> mdist m2 ((f \<circ> \<rho>) n) ((f \<circ> \<rho>) n') < \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2793
        using f\<rho> unfolding L.M2.MCauchy_def by (meson half_gt_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2794
      have "\<forall>\<^sub>F n in sequentially. inverse (Suc n) < \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2795
        using Archimedean_eventually_inverse \<open>0 < \<epsilon>\<close> divide_pos_pos zero_less_numeral by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2796
      then obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> inverse (Suc n) < \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2797
        by (meson eventually_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2798
      have "mdist m2 ((f \<circ> \<sigma>) n) ((f \<circ> \<sigma>) n') < \<epsilon>" if "n \<ge> max M N" "n' \<ge> max M N" for n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2799
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2800
        have "mdist m2 ((f \<circ> \<sigma>) n) ((f \<circ> \<sigma>) n') \<le> mdist m2 ((f \<circ> \<sigma>) n) ((f \<circ> \<rho>) n) + mdist m2 ((f \<circ> \<rho>) n) ((f \<circ> \<sigma>) n')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2801
          by (meson L.M2.MCauchy_def \<open>range (f \<circ> \<sigma>) \<subseteq> mspace m2\<close> f\<rho> mdist_triangle rangeI subset_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2802
        also have "\<dots> \<le> mdist m2 ((f \<circ> \<sigma>) n) ((f \<circ> \<rho>) n) + mdist m2 ((f \<circ> \<rho>) n) ((f \<circ> \<rho>) n') + mdist m2 ((f \<circ> \<rho>) n') ((f \<circ> \<sigma>) n')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2803
          by (smt (verit) L.M2.MCauchy_def L.M2.triangle \<open>range (f \<circ> \<sigma>) \<subseteq> mspace m2\<close> f\<rho> range_subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2804
        also have "\<dots> < \<epsilon>/4 + \<epsilon>/2 + \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2805
          using \<rho>2[of n] \<rho>2[of n'] N[of n] N[of n'] that M[of n n'] by (simp add: L.M2.commute)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2806
        also have "... \<le> \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2807
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2808
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2809
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2810
      then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> mdist m2 ((f \<circ> \<sigma>) n) ((f \<circ> \<sigma>) n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2811
        by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2812
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2813
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2814
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2815
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2816
lemma Cauchy_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2817
  assumes "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2818
    and "S \<subseteq> T" and "T \<subseteq> (mtopology_of m1) closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2819
    and "continuous_map (subtopology (mtopology_of m1) T) (mtopology_of m2) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2820
  shows "Cauchy_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2821
  by (metis Cauchy_continuous_map_on_intermediate_closure_aux assms closure_of_subset_topspace order.trans topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2822
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2823
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2824
lemma Cauchy_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2825
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2826
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2827
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2828
  where "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2829
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2830
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2831
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2832
    where g: "continuous_map (subtopology (mtopology_of m1) ((mtopology_of m1) closure_of S)) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2833
                        (mtopology_of m2) g"  "(\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2834
    by (metis Cauchy_continuous_map_extends_to_continuous_closure_of f m2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2835
  have "Cauchy_continuous_map (submetric m1 (mtopology_of m1 closure_of S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2836
  proof (rule Cauchy_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2837
    show "Cauchy_continuous_map (submetric m1 (mspace m1 \<inter> S)) m2 g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2838
      by (smt (verit, best) IntD2 Cauchy_continuous_map_eq f g(2) inf_commute mspace_submetric submetric_restrict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2839
    show "mspace m1 \<inter> S \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2840
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2841
    show "mtopology_of m1 closure_of S \<subseteq> mtopology_of m1 closure_of (mspace m1 \<inter> S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2842
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2843
    show "continuous_map (subtopology (mtopology_of m1) (mtopology_of m1 closure_of S)) (mtopology_of m2) g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2844
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2845
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2846
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2847
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2848
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2849
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2850
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2851
lemma Cauchy_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2852
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2853
    and "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2854
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2855
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2856
  where "Cauchy_continuous_map (submetric m1 T) m2 g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2857
  by (metis Cauchy_continuous_map_extends_to_closure_of Cauchy_continuous_map_from_submetric_mono assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2858
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2859
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2860
subsection\<open>Metric space of bounded functions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2861
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2862
context Metric_space
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2863
begin
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2864
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2865
definition fspace :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) set" where 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2866
  "fspace \<equiv> \<lambda>S. {f. f`S \<subseteq> M \<and> f \<in> extensional S \<and> mbounded (f`S)}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2867
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2868
definition fdist :: "['b set, 'b \<Rightarrow> 'a, 'b \<Rightarrow> 'a] \<Rightarrow> real" where 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2869
  "fdist \<equiv> \<lambda>S f g. if f \<in> fspace S \<and> g \<in> fspace S \<and> S \<noteq> {} 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2870
                    then Sup ((\<lambda>x. d (f x) (g x)) ` S) else 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2871
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2872
lemma fspace_empty [simp]: "fspace {} = {\<lambda>x. undefined}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2873
  by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2874
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2875
lemma fdist_empty [simp]: "fdist {} = (\<lambda>x y. 0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2876
  by (auto simp: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2877
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2878
lemma fspace_in_M: "\<lbrakk>f \<in> fspace S; x \<in> S\<rbrakk> \<Longrightarrow> f x \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2879
  by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2880
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2881
lemma bdd_above_dist:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2882
  assumes f: "f \<in> fspace S" and g: "g \<in> fspace S" and "S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2883
  shows "bdd_above ((\<lambda>u. d (f u) (g u)) ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2884
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2885
  obtain a where "a \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2886
    using \<open>S \<noteq> {}\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2887
  obtain B x C y where "B>0"  and B: "f`S \<subseteq> mcball x B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2888
    and "C>0" and C: "g`S \<subseteq> mcball y C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2889
    using f g mbounded_pos by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2890
  have "d (f u) (g u) \<le> B + d x y + C" if "u\<in>S" for u 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2891
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2892
    have "f u \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2893
      by (meson B image_eqI mbounded_mcball mbounded_subset_mspace subsetD that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2894
    have "g u \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2895
      by (meson C image_eqI mbounded_mcball mbounded_subset_mspace subsetD that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2896
    have "x \<in> M" "y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2897
      using B C that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2898
    have "d (f u) (g u) \<le> d (f u) x + d x (g u)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2899
      by (simp add: \<open>f u \<in> M\<close> \<open>g u \<in> M\<close> \<open>x \<in> M\<close> triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2900
    also have "\<dots> \<le> d (f u) x + d x y + d y (g u)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2901
      by (simp add: \<open>f u \<in> M\<close> \<open>g u \<in> M\<close> \<open>x \<in> M\<close> \<open>y \<in> M\<close> triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2902
    also have "\<dots> \<le> B + d x y + C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2903
      using B C commute that by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2904
    finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2905
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2906
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2907
    by (meson bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2908
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2909
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2910
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2911
lemma Metric_space_funspace: "Metric_space (fspace S) (fdist S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2912
proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2913
  show *: "0 \<le> fdist S f g" for f g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2914
    by (auto simp: fdist_def intro: cSUP_upper2 [OF bdd_above_dist])
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2915
  show "fdist S f g = fdist S g f" for f g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2916
    by (auto simp: fdist_def commute)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2917
  show "(fdist S f g = 0) = (f = g)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2918
    if fg: "f \<in> fspace S" "g \<in> fspace S" for f g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2919
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2920
    assume 0: "fdist S f g = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2921
    show "f = g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2922
    proof (cases "S={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2923
      case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2924
      with 0 that show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2925
        by (simp add: fdist_def fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2926
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2927
      case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2928
      with 0 fg have Sup0: "(SUP x\<in>S. d (f x) (g x)) = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2929
        by (simp add: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2930
      have "d (f x) (g x) = 0" if "x\<in>S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2931
          by (smt (verit) False Sup0 \<open>x\<in>S\<close> bdd_above_dist [OF fg] less_cSUP_iff nonneg)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2932
      with fg show "f=g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2933
        by (simp add: fspace_def extensionalityI image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2934
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2935
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2936
    show "f = g \<Longrightarrow> fdist S f g = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2937
      using fspace_in_M [OF \<open>g \<in> fspace S\<close>] by (auto simp: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2938
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2939
  show "fdist S f h \<le> fdist S f g + fdist S g h"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2940
    if fgh: "f \<in> fspace S" "g \<in> fspace S" "h \<in> fspace S" for f g h
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2941
  proof (clarsimp simp add: fdist_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2942
    assume "S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2943
    have dfh: "d (f x) (h x) \<le> d (f x) (g x) + d (g x) (h x)" if "x\<in>S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2944
      by (meson fgh fspace_in_M that triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2945
    have bdd_fgh: "bdd_above ((\<lambda>x. d (f x) (g x)) ` S)" "bdd_above ((\<lambda>x. d (g x) (h x)) ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2946
      by (simp_all add: \<open>S \<noteq> {}\<close> bdd_above_dist that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2947
    then obtain B C where B: "\<And>x. x\<in>S \<Longrightarrow> d (f x) (g x) \<le> B" and C: "\<And>x. x\<in>S \<Longrightarrow> d (g x) (h x) \<le> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2948
      by (auto simp: bdd_above_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2949
    then have "\<And>x. x\<in>S \<Longrightarrow> d (f x) (g x) + d (g x) (h x) \<le> B+C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2950
      by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2951
    then have bdd: "bdd_above ((\<lambda>x. d (f x) (g x) + d (g x) (h x)) ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2952
      by (auto simp: bdd_above_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2953
    then have "(SUP x\<in>S. d (f x) (h x)) \<le> (SUP x\<in>S. d (f x) (g x) + d (g x) (h x))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2954
      by (metis (mono_tags, lifting) cSUP_mono \<open>S \<noteq> {}\<close> dfh)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2955
    also have "\<dots> \<le> (SUP x\<in>S. d (f x) (g x)) + (SUP x\<in>S. d (g x) (h x))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2956
      by (simp add: \<open>S \<noteq> {}\<close> bdd cSUP_le_iff bdd_fgh add_mono cSup_upper)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2957
    finally show "(SUP x\<in>S. d (f x) (h x)) \<le> (SUP x\<in>S. d (f x) (g x)) + (SUP x\<in>S. d (g x) (h x))" .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2958
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2959
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2960
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2961
end
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2962
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2963
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2964
definition funspace where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2965
  "funspace S m \<equiv> metric (Metric_space.fspace (mspace m) (mdist m) S, 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2966
                          Metric_space.fdist (mspace m) (mdist m) S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2967
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2968
lemma mspace_funspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2969
  "mspace (funspace S m) = Metric_space.fspace (mspace m) (mdist m) S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2970
  by (simp add: Metric_space.Metric_space_funspace Metric_space.mspace_metric funspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2971
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2972
lemma mdist_funspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2973
  "mdist (funspace S m) = Metric_space.fdist (mspace m) (mdist m) S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2974
  by (simp add: Metric_space.Metric_space_funspace Metric_space.mdist_metric funspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2975
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2976
lemma funspace_imp_welldefined:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2977
   "\<lbrakk>f \<in> mspace (funspace S m); x \<in> S\<rbrakk> \<Longrightarrow> f x \<in> mspace m"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2978
  by (simp add: Metric_space.fspace_def subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2979
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2980
lemma funspace_imp_extensional:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2981
   "f \<in> mspace (funspace S m) \<Longrightarrow> f \<in> extensional S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2982
  by (simp add: Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2983
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2984
lemma funspace_imp_bounded_image:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2985
   "f \<in> mspace (funspace S m) \<Longrightarrow> Metric_space.mbounded (mspace m) (mdist m) (f ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2986
  by (simp add: Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2987
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2988
lemma funspace_imp_bounded:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2989
   "f \<in> mspace (funspace S m) \<Longrightarrow> S = {} \<or> (\<exists>c B. \<forall>x \<in> S. mdist m c (f x) \<le> B)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2990
  by (auto simp: Metric_space.fspace_def Metric_space.mbounded)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2991
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2992
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2993
lemma (in Metric_space) funspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2994
  assumes "f \<in> fspace S" "g \<in> fspace S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2995
  obtains B where "\<And>x. x \<in> S \<Longrightarrow> d (f x) (g x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2996
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2997
  have "mbounded (f ` S \<union> g ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2998
    using mbounded_Un assms by (force simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2999
  then show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3000
    by (metis UnCI imageI mbounded_alt that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3001
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3002
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3003
lemma funspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3004
  assumes "f \<in> mspace (funspace S m)" "g \<in> mspace (funspace S m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3005
  obtains B where "\<And>x. x \<in> S \<Longrightarrow> mdist m (f x) (g x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3006
  by (metis Metric_space_mspace_mdist assms mspace_funspace Metric_space.funspace_imp_bounded2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3007
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3008
lemma (in Metric_space) funspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3009
  assumes fg: "f \<in> fspace S" "g \<in> fspace S" and "S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3010
  shows "fdist S f g \<le> a \<longleftrightarrow> (\<forall>x \<in> S. d (f x) (g x) \<le> a)" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3011
    using assms bdd_above_dist [OF fg] by (simp add: fdist_def cSUP_le_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3012
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3013
lemma funspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3014
  assumes "f \<in> mspace (funspace S m)" "g \<in> mspace (funspace S m)" and "S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3015
  shows "mdist (funspace S m) f g \<le> a \<longleftrightarrow> (\<forall>x \<in> S. mdist m (f x) (g x) \<le> a)" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3016
  using assms by (simp add: Metric_space.funspace_mdist_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3017
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3018
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3019
lemma (in Metric_space) mcomplete_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3020
  assumes "mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3021
  shows "mcomplete_of (funspace S Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3022
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3023
  interpret F: Metric_space "fspace S" "fdist S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3024
    by (simp add: Metric_space_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3025
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3026
  proof (cases "S={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3027
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3028
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3029
      by (simp add: mcomplete_of_def mcomplete_trivial_singleton)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3030
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3031
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3032
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3033
    proof (clarsimp simp: mcomplete_of_def Metric_space.mcomplete_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3034
      fix \<sigma>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3035
      assume \<sigma>: "F.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3036
      then have \<sigma>M: "\<And>n x. x \<in> S \<Longrightarrow> \<sigma> n x \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3037
        by (auto simp: F.MCauchy_def intro: fspace_in_M)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3038
      have fdist_less: "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> fdist S (\<sigma> n) (\<sigma> n') < \<epsilon>" if "\<epsilon>>0" for \<epsilon>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3039
        using \<sigma> that by (auto simp: F.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3040
      have \<sigma>ext: "\<And>n. \<sigma> n \<in> extensional S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3041
        using \<sigma> unfolding F.MCauchy_def by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3042
      have \<sigma>bd: "\<And>n. mbounded (\<sigma> n ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3043
        using \<sigma> unfolding F.MCauchy_def by (simp add: fspace_def image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3044
      have \<sigma>in[simp]: "\<sigma> n \<in> fspace S" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3045
        using F.MCauchy_def \<sigma> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3046
      have bd2: "\<And>n n'. \<exists>B. \<forall>x \<in> S. d (\<sigma> n x) (\<sigma> n' x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3047
        using \<sigma> unfolding F.MCauchy_def by (metis range_subsetD funspace_imp_bounded2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3048
      have sup: "\<And>n n' x0. x0 \<in> S \<Longrightarrow> d (\<sigma> n x0) (\<sigma> n' x0) \<le> Sup ((\<lambda>x. d (\<sigma> n x) (\<sigma> n' x)) ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3049
      proof (rule cSup_upper)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3050
        show "bdd_above ((\<lambda>x. d (\<sigma> n x) (\<sigma> n' x)) ` S)" if "x0 \<in> S" for n n' x0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3051
          using that bd2 by (meson bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3052
      qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3053
      have pcy: "MCauchy (\<lambda>n. \<sigma> n x)" if "x \<in> S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3054
        unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3055
      proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3056
        show "range (\<lambda>n. \<sigma> n x) \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3057
          using \<sigma>M that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3058
        fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3059
        assume "\<epsilon> > 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3060
        then obtain N where N: "\<And>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> fdist S (\<sigma> n) (\<sigma> n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3061
          using \<sigma> by (force simp: F.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3062
        { fix n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3063
          assume n: "N \<le> n" "N \<le> n'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3064
          have "d (\<sigma> n x) (\<sigma> n' x) \<le> (SUP x\<in>S. d (\<sigma> n x) (\<sigma> n' x))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3065
            using that sup by presburger
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3066
          then have "d (\<sigma> n x) (\<sigma> n' x) \<le> fdist S (\<sigma> n) (\<sigma> n')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3067
            by (simp add: fdist_def \<open>S \<noteq> {}\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3068
          with N n have "d (\<sigma> n x) (\<sigma> n' x) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3069
            by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3070
        } then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n x) (\<sigma> n' x) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3071
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3072
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3073
      have "\<exists>l. limitin mtopology (\<lambda>n. \<sigma> n x) l sequentially" if "x \<in> S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3074
        using assms mcomplete_def pcy \<open>x \<in> S\<close> by presburger
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3075
      then obtain g0 where g0: "\<And>x. x \<in> S \<Longrightarrow> limitin mtopology (\<lambda>n. \<sigma> n x) (g0 x) sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3076
        by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3077
      define g where "g \<equiv> restrict g0 S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3078
      have gext: "g \<in> extensional S" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3079
       and glim: "\<And>x. x \<in> S \<Longrightarrow> limitin mtopology (\<lambda>n. \<sigma> n x) (g x) sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3080
        by (auto simp: g_def g0)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3081
      have gwd: "g x \<in> M" if "x \<in> S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3082
        using glim limitin_metric that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3083
      have unif: "\<exists>N. \<forall>x n. x \<in> S \<longrightarrow> N \<le> n \<longrightarrow> d (\<sigma> n x) (g x) < \<epsilon>" if "\<epsilon>>0" for \<epsilon>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3084
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3085
        obtain N where N: "\<And>n n'. N \<le> n \<and> N \<le> n' \<Longrightarrow> Sup ((\<lambda>x. d (\<sigma> n x) (\<sigma> n' x)) ` S) < \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3086
          using \<open>S\<noteq>{}\<close> \<open>\<epsilon>>0\<close> fdist_less [of "\<epsilon>/2"]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3087
          by (metis (mono_tags) \<sigma>in fdist_def half_gt_zero) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3088
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3089
        proof (intro exI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3090
          fix x n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3091
          assume "x \<in> S" and "N \<le> n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3092
          obtain N' where N': "\<And>n. N' \<le> n \<Longrightarrow> \<sigma> n x \<in> M \<and> d (\<sigma> n x) (g x) < \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3093
            by (metis \<open>0 < \<epsilon>\<close> \<open>x \<in> S\<close> glim half_gt_zero limit_metric_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3094
          have "d (\<sigma> n x) (g x) \<le> d (\<sigma> n x) (\<sigma> (max N N') x) + d (\<sigma> (max N N') x) (g x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3095
            using \<open>x \<in> S\<close> \<sigma>M gwd triangle by presburger
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3096
          also have "\<dots> < \<epsilon>/2 + \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3097
            by (smt (verit) N N' \<open>N \<le> n\<close> \<open>x \<in> S\<close> max.cobounded1 max.cobounded2 sup)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3098
          finally show "d (\<sigma> n x) (g x) < \<epsilon>" by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3099
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3100
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3101
      have "limitin F.mtopology \<sigma> g sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3102
        unfolding F.limit_metric_sequentially
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3103
      proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3104
        obtain N where N: "\<And>n n'. N \<le> n \<and> N \<le> n' \<Longrightarrow> Sup ((\<lambda>x. d (\<sigma> n x) (\<sigma> n' x)) ` S) < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3105
          using fdist_less [of 1] \<open>S\<noteq>{}\<close> by (auto simp: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3106
        have "\<And>x. x \<in> \<sigma> N ` S \<Longrightarrow> x \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3107
          using \<sigma>M by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3108
        obtain a B where "a \<in> M" and B: "\<And>x. x \<in> (\<sigma> N) ` S \<Longrightarrow> d a x \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3109
          by (metis False \<sigma>M \<sigma>bd ex_in_conv imageI mbounded_alt_pos)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3110
        have "d a (g x) \<le> B+1" if "x\<in>S" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3111
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3112
          have "d a (g x) \<le> d a (\<sigma> N x) + d (\<sigma> N x) (g x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3113
            by (simp add: \<open>a \<in> M\<close> \<sigma>M gwd that triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3114
          also have "\<dots> \<le> B+1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3115
          proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3116
            have "d a (\<sigma> N x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3117
              by (simp add: B that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3118
            moreover 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3119
            have False if 1: "d (\<sigma> N x) (g x) > 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3120
            proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3121
              obtain r where "1 < r" and r: "r < d (\<sigma> N x) (g x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3122
                using 1 dense by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3123
              then obtain N' where N': "\<And>n. N' \<le> n \<Longrightarrow> \<sigma> n x \<in> M \<and> d (\<sigma> n x) (g x) < r-1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3124
                using glim [OF \<open>x\<in>S\<close>] by (fastforce simp: limit_metric_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3125
              have "d (\<sigma> N x) (g x) \<le> d (\<sigma> N x) (\<sigma> (max N N') x) + d (\<sigma> (max N N') x) (g x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3126
                by (metis \<open>x \<in> S\<close> \<sigma>M commute gwd triangle')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3127
              also have "\<dots> < 1 + (r-1)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3128
                by (smt (verit) N N' \<open>x \<in> S\<close> max.cobounded1 max.cobounded2 max.idem sup)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3129
              finally have "d (\<sigma> N x) (g x) < r"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3130
                by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3131
              with r show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3132
                by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3133
            qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3134
            ultimately show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3135
              by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3136
          qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3137
          finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3138
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3139
        with gwd \<open>a \<in> M\<close> have "mbounded (g ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3140
          unfolding mbounded by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3141
        with gwd gext show "g \<in> fspace S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3142
          by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3143
        fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3144
        assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3145
        then obtain N where "\<And>x n. x \<in> S \<Longrightarrow> N \<le> n \<Longrightarrow> d (\<sigma> n x) (g x) < \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3146
          by (meson unif half_gt_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3147
        then have "fdist S (\<sigma> n) g \<le> \<epsilon>/2" if "N \<le> n" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3148
          using \<open>g \<in> fspace S\<close> False that
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3149
          by (force simp: funspace_mdist_le simp del: divide_const_simps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3150
        then show "\<exists>N. \<forall>n\<ge>N. \<sigma> n \<in> fspace S \<and> fdist S (\<sigma> n) g < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3151
          by (metis \<open>0 < \<epsilon>\<close> \<sigma>in add_strict_increasing field_sum_of_halves half_gt_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3152
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3153
      then show "\<exists>x. limitin F.mtopology \<sigma> x sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3154
        by blast 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3155
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3156
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3157
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3158
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3159
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3160
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3161
subsection\<open>Metric space of continuous bounded functions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3162
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3163
definition cfunspace where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3164
  "cfunspace X m \<equiv> submetric (funspace (topspace X) m) {f. continuous_map X (mtopology_of m) f}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3165
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3166
lemma mspace_cfunspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3167
  "mspace (cfunspace X m) = 
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3168
     {f. f \<in> topspace X \<rightarrow> mspace m \<and> f \<in> extensional (topspace X) \<and>
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3169
         Metric_space.mbounded (mspace m) (mdist m) (f ` (topspace X)) \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3170
         continuous_map X (mtopology_of m) f}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3171
  by (auto simp: cfunspace_def Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3172
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3173
lemma mdist_cfunspace_eq_mdist_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3174
  "mdist (cfunspace X m) = mdist (funspace (topspace X) m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3175
  by (auto simp: cfunspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3176
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3177
lemma cfunspace_subset_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3178
   "mspace (cfunspace X m) \<subseteq> mspace (funspace (topspace X) m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3179
  by (simp add: cfunspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3180
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3181
lemma cfunspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3182
   "\<lbrakk>f \<in> mspace (cfunspace X m); g \<in> mspace (cfunspace X m); topspace X \<noteq> {}\<rbrakk>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3183
     \<Longrightarrow> mdist (cfunspace X m) f g \<le> a \<longleftrightarrow> (\<forall>x \<in> topspace X. mdist m (f x) (g x) \<le> a)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3184
  by (simp add: cfunspace_def Metric_space.funspace_mdist_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3185
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3186
lemma cfunspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3187
  assumes "f \<in> mspace (cfunspace X m)" "g \<in> mspace (cfunspace X m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3188
  obtains B where "\<And>x. x \<in> topspace X \<Longrightarrow> mdist m (f x) (g x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3189
  by (metis assms all_not_in_conv cfunspace_mdist_le nle_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3190
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3191
lemma cfunspace_mdist_lt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3192
   "\<lbrakk>compactin X (topspace X); f \<in> mspace (cfunspace X m);
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3193
     g \<in> mspace (cfunspace X m); mdist (cfunspace X m) f g < a;
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3194
     x \<in> topspace X\<rbrakk>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3195
     \<Longrightarrow> mdist m (f x) (g x) < a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3196
  by (metis (full_types) cfunspace_mdist_le empty_iff less_eq_real_def less_le_not_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3197
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3198
lemma mdist_cfunspace_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3199
  assumes "0 \<le> B" and B: "\<And>x. x \<in> topspace X \<Longrightarrow> mdist m (f x) (g x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3200
  shows "mdist (cfunspace X m) f g \<le> B"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3201
proof (cases "X = trivial_topology")
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3202
  case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3203
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3204
    by (simp add: Metric_space.fdist_empty \<open>B\<ge>0\<close> cfunspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3205
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3206
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3207
  have bdd: "bdd_above ((\<lambda>u. mdist m (f u) (g u)) ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3208
    by (meson B bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3209
  with assms bdd show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3210
    by (simp add: mdist_cfunspace_eq_mdist_funspace Metric_space.fdist_def cSUP_le_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3211
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3212
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3213
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3214
lemma mdist_cfunspace_imp_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3215
   "\<lbrakk>f \<in> mspace (cfunspace X m); g \<in> mspace (cfunspace X m);
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3216
     mdist (cfunspace X m) f g \<le> a; x \<in> topspace X\<rbrakk> \<Longrightarrow> mdist m (f x) (g x) \<le> a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3217
  using cfunspace_mdist_le by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3218
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3219
lemma compactin_mspace_cfunspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3220
   "compactin X (topspace X)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3221
     \<Longrightarrow> mspace (cfunspace X m) =
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3222
          {f. (\<forall>x \<in> topspace X. f x \<in> mspace m) \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3223
               f \<in> extensional (topspace X) \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3224
               continuous_map X (mtopology_of m) f}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3225
  by (auto simp: Metric_space.compactin_imp_mbounded image_compactin mtopology_of_def) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3226
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3227
lemma (in Metric_space) mcomplete_cfunspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3228
  assumes "mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3229
  shows "mcomplete_of (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3230
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3231
  interpret F: Metric_space "fspace (topspace X)" "fdist (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3232
    by (simp add: Metric_space_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3233
  interpret S: Submetric "fspace (topspace X)" "fdist (topspace X)" "mspace (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3234
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3235
    show "mspace (cfunspace X Self) \<subseteq> fspace (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3236
      by (metis cfunspace_subset_funspace mdist_Self mspace_Self mspace_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3237
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3238
  show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3239
  proof (cases "X = trivial_topology")
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3240
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3241
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3242
      by (simp add: mcomplete_of_def mcomplete_trivial_singleton mdist_cfunspace_eq_mdist_funspace cong: conj_cong)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3243
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3244
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3245
    have *: "continuous_map X mtopology g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3246
      if "range \<sigma> \<subseteq> mspace (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3247
        and g: "limitin F.mtopology \<sigma> g sequentially" for \<sigma> g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3248
      unfolding continuous_map_to_metric
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3249
    proof (intro strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3250
      have \<sigma>: "\<And>n. continuous_map X mtopology (\<sigma> n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3251
        using that by (auto simp: mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3252
      fix x and \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3253
      assume "x \<in> topspace X" and "0 < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3254
      then obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<sigma> n \<in> fspace (topspace X) \<and> fdist (topspace X) (\<sigma> n) g < \<epsilon>/3"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3255
        unfolding mtopology_of_def F.limitin_metric
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3256
        by (metis F.limit_metric_sequentially divide_pos_pos g zero_less_numeral) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3257
      then obtain U where "openin X U" "x \<in> U" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3258
        and U: "\<And>y. y \<in> U \<Longrightarrow> \<sigma> N y \<in> mball (\<sigma> N x) (\<epsilon>/3)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3259
        by (metis Metric_space.continuous_map_to_metric Metric_space_axioms \<open>0 < \<epsilon>\<close> \<open>x \<in> topspace X\<close> \<sigma> divide_pos_pos zero_less_numeral)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3260
      moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3261
      have "g y \<in> mball (g x) \<epsilon>" if "y\<in>U" for y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3262
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3263
        have "U \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3264
          using \<open>openin X U\<close> by (simp add: openin_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3265
        have gx: "g x \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3266
          by (meson F.limitin_mspace \<open>x \<in> topspace X\<close> fspace_in_M g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3267
        have "y \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3268
          using \<open>U \<subseteq> topspace X\<close> that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3269
        have gy: "g y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3270
          by (meson F.limitin_mspace[OF g] \<open>U \<subseteq> topspace X\<close> fspace_in_M subsetD that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3271
        have "d (g x) (g y) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3272
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3273
          have *: "d (\<sigma> N x0) (g x0) \<le> \<epsilon>/3" if "x0 \<in> topspace X" for x0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3274
          proof -
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3275
            have "g \<in> fspace (topspace X)"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3276
              using F.limit_metric_sequentially g by blast
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3277
            with N that have "bdd_above ((\<lambda>x. d (\<sigma> N x) (g x)) ` topspace X)"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3278
              by (force intro: bdd_above_dist)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3279
            then have "d (\<sigma> N x0) (g x0) \<le> Sup ((\<lambda>x. d (\<sigma> N x) (g x)) ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3280
              by (simp add: cSup_upper that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3281
            also have "\<dots> \<le> \<epsilon>/3"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3282
              using g False N \<open>g \<in> fspace (topspace X)\<close> 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  3283
                by (fastforce simp: F.limit_metric_sequentially fdist_def)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3284
            finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3285
          qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3286
          have "d (g x) (g y) \<le> d (g x) (\<sigma> N x) + d (\<sigma> N x) (g y)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3287
            using U gx gy that triangle by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3288
          also have "\<dots> < \<epsilon>/3 + \<epsilon>/3 + \<epsilon>/3"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3289
            by (smt (verit) "*" U gy \<open>x \<in> topspace X\<close> \<open>y \<in> topspace X\<close> commute in_mball that triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3290
          finally show ?thesis by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3291
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3292
        with gx gy show ?thesis by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3293
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3294
      ultimately show "\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y\<in>U. g y \<in> mball (g x) \<epsilon>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3295
        by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3296
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3297
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3298
    have "S.sub.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3299
    proof (rule S.sequentially_closedin_mcomplete_imp_mcomplete)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3300
      show "F.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3301
        by (metis assms mcomplete_funspace mcomplete_of_def mdist_Self mdist_funspace mspace_Self mspace_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3302
      fix \<sigma> g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3303
      assume g: "range \<sigma> \<subseteq> mspace (cfunspace X Self) \<and> limitin F.mtopology \<sigma> g sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3304
      show "g \<in> mspace (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3305
      proof (simp add: mtopology_of_def, intro conjI)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3306
        show "g \<in> topspace X \<rightarrow> M" "g \<in> extensional (topspace X)" "mbounded (g ` topspace X)"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3307
          using g F.limitin_mspace by (force simp: fspace_def)+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3308
        show "continuous_map X mtopology g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3309
          using * g by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3310
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3311
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3312
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3313
      by (simp add: mcomplete_of_def mdist_cfunspace_eq_mdist_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3314
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3315
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3316
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3317
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3318
subsection\<open>Existence of completion for any metric space M as a subspace of M=>R\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3319
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3320
lemma (in Metric_space) metric_completion_explicit:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3321
  obtains f :: "['a,'a] \<Rightarrow> real" and S where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3322
      "S \<subseteq> mspace(funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3323
      "mcomplete_of (submetric (funspace M euclidean_metric) S)"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3324
      "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3325
      "mtopology_of(funspace M euclidean_metric) closure_of f ` M = S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3326
      "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3327
            \<Longrightarrow> mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3328
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3329
  define m':: "('a\<Rightarrow>real) metric" where "m' \<equiv> funspace M euclidean_metric"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3330
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3331
  proof (cases "M = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3332
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3333
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3334
      using that by (simp add: mcomplete_of_def mcomplete_trivial)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3335
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3336
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3337
    then obtain a where "a \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3338
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3339
    define f where "f \<equiv> \<lambda>x. (\<lambda>u \<in> M. d x u - d a u)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3340
    define S where "S \<equiv> mtopology_of(funspace M euclidean_metric) closure_of (f ` M)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3341
    interpret S: Submetric "Met_TC.fspace M" "Met_TC.fdist M" "S \<inter> Met_TC.fspace M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3342
      by (simp add: Met_TC.Metric_space_funspace Submetric.intro Submetric_axioms_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3343
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3344
    have fim: "f ` M \<subseteq> mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3345
    proof (clarsimp simp: m'_def Met_TC.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3346
      fix b
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3347
      assume "b \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3348
      then have "\<And>c. \<lbrakk>c \<in> M\<rbrakk> \<Longrightarrow> \<bar>d b c - d a c\<bar> \<le> d a b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3349
        by (smt (verit, best) \<open>a \<in> M\<close> commute triangle'')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3350
      then have "(\<lambda>x. d b x - d a x) ` M \<subseteq> cball 0 (d a b)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3351
        by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3352
      then show "f b \<in> extensional M \<and> bounded (f b ` M)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3353
        by (metis bounded_cball bounded_subset f_def image_restrict_eq restrict_extensional_sub set_eq_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3354
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3355
    show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3356
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3357
      show "S \<subseteq> mspace (funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3358
        by (simp add: S_def in_closure_of subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3359
      have "closedin S.mtopology (S \<inter> Met_TC.fspace M)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3360
        by (simp add: S_def closedin_Int funspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3361
      moreover have "S.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3362
        using Metric_space.mcomplete_funspace Met_TC.Metric_space_axioms by (fastforce simp: mcomplete_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3363
      ultimately show "mcomplete_of (submetric (funspace M euclidean_metric) S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3364
        by (simp add: S.closedin_eq_mcomplete mcomplete_of_def)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3365
      show "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3366
        using S_def fim in_closure_of m'_def by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3367
      show "mtopology_of (funspace M euclidean_metric) closure_of f ` M = S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3368
        by (auto simp: f_def S_def mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3369
      show "mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3370
        if "x \<in> M" "y \<in> M" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3371
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3372
        have "\<forall>c\<in>M. dist (f x c) (f y c) \<le> r \<Longrightarrow> d x y \<le> r" for r
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3373
          using that by (auto simp: f_def dist_real_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3374
        moreover have "dist (f x z) (f y z) \<le> r" if "d x y \<le> r" and "z \<in> M" for r z
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3375
          using that \<open>x \<in> M\<close> \<open>y \<in> M\<close>  
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3376
          apply (simp add: f_def Met_TC.fdist_def dist_real_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3377
          by (smt (verit, best) commute triangle')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3378
        ultimately have "(SUP c \<in> M. dist (f x c) (f y c)) = d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3379
          by (intro cSup_unique) auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3380
        with that fim show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3381
          using that fim by (simp add: Met_TC.fdist_def False m'_def image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3382
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3383
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3384
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3385
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3386
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3387
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3388
lemma (in Metric_space) metric_completion:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3389
  obtains f :: "['a,'a] \<Rightarrow> real" and m' where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3390
    "mcomplete_of m'"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3391
    "f \<in> M \<rightarrow> mspace m' "
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3392
    "mtopology_of m' closure_of f ` M = mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3393
    "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> mdist m' (f x) (f y) = d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3394
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3395
  obtain f :: "['a,'a] \<Rightarrow> real" and S where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3396
    Ssub: "S \<subseteq> mspace(funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3397
    and mcom: "mcomplete_of (submetric (funspace M euclidean_metric) S)"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3398
    and fim: "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3399
    and eqS: "mtopology_of(funspace M euclidean_metric) closure_of f ` M = S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3400
    and eqd: "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> mdist (funspace M euclidean_metric) (f x) (f y) = d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3401
    using metric_completion_explicit by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3402
  define m' where "m' \<equiv> submetric (funspace M euclidean_metric) S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3403
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3404
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3405
    show "mcomplete_of m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3406
      by (simp add: mcom m'_def)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3407
    show "f \<in> M \<rightarrow> mspace m'"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3408
      using Ssub fim m'_def by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3409
    show "mtopology_of m' closure_of f ` M = mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3410
      using eqS fim Ssub
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3411
      by (force simp: m'_def mtopology_of_submetric closure_of_subtopology Int_absorb1 image_subset_iff_funcset)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3412
    show "mdist m' (f x) (f y) = d x y" if "x \<in> M" and "y \<in> M" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3413
      using that eqd m'_def by force 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3414
  qed 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3415
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3416
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3417
lemma metrizable_space_completion:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3418
  assumes "metrizable_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3419
  obtains f :: "['a,'a] \<Rightarrow> real" and Y where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3420
       "completely_metrizable_space Y" "embedding_map X Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3421
                "Y closure_of (f ` (topspace X)) = topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3422
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3423
  obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3424
    using assms metrizable_space_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3425
  then interpret Metric_space M d by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3426
  obtain f :: "['a,'a] \<Rightarrow> real" and m' where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3427
    "mcomplete_of m'"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3428
    and fim: "f \<in> M \<rightarrow> mspace m' "
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3429
    and m': "mtopology_of m' closure_of f ` M = mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3430
    and eqd: "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> mdist m' (f x) (f y) = d x y"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3431
    by (metis metric_completion)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3432
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3433
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3434
    show "completely_metrizable_space (mtopology_of m')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3435
      using \<open>mcomplete_of m'\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3436
      unfolding completely_metrizable_space_def mcomplete_of_def mtopology_of_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3437
      by (metis Metric_space_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3438
    show "embedding_map X (mtopology_of m') f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3439
      using Metric_space12.isometry_imp_embedding_map
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3440
      by (metis Metric_space12_def Metric_space_axioms Metric_space_mspace_mdist Xeq eqd fim 
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3441
             mtopology_of_def)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3442
    show "(mtopology_of m') closure_of f ` topspace X = topspace (mtopology_of m')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3443
      by (simp add: Xeq m')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3444
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3445
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3446
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3447
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3448
subsection\<open>Contractions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3449
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3450
lemma (in Metric_space) contraction_imp_unique_fixpoint:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3451
  assumes "f x = x" "f y = y"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3452
    and "f \<in> M \<rightarrow> M"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3453
    and "k < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3454
    and "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> d (f x) (f y) \<le> k * d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3455
    and "x \<in> M" "y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3456
  shows "x = y"
78283
6fa0812135e0 more small simplifications
paulson <lp15@cam.ac.uk>
parents: 78274
diff changeset
  3457
  by (smt (verit, ccfv_SIG) mdist_pos_less mult_le_cancel_right1 assms)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3458
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3459
text \<open>Banach Fixed-Point Theorem (aka, Contraction Mapping Principle)\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3460
lemma (in Metric_space) Banach_fixedpoint_thm:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3461
  assumes mcomplete and "M \<noteq> {}" and fim: "f \<in> M \<rightarrow> M"    
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3462
    and "k < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3463
    and con: "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> d (f x) (f y) \<le> k * d x y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3464
  obtains x where "x \<in> M" "f x = x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3465
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3466
  obtain a where "a \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3467
    using \<open>M \<noteq> {}\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3468
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3469
  proof (cases "\<forall>x \<in> M. f x = f a")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3470
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3471
    then show ?thesis
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3472
      by (metis \<open>a \<in> M\<close> fim image_subset_iff image_subset_iff_funcset that)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3473
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3474
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3475
    then obtain b where "b \<in> M" and b: "f b \<noteq> f a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3476
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3477
    have "k>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3478
      using Lipschitz_coefficient_pos [where f=f]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3479
      by (metis False \<open>a \<in> M\<close> con fim mdist_Self mspace_Self)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3480
    define \<sigma> where "\<sigma> \<equiv> \<lambda>n. (f^^n) a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3481
    have f_iter: "\<sigma> n \<in> M" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3482
      unfolding \<sigma>_def by (induction n) (use \<open>a \<in> M\<close> fim in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3483
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3484
    proof (cases "f a = a")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3485
      case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3486
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3487
        using \<open>a \<in> M\<close> that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3488
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3489
      case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3490
      have "MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3491
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3492
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3493
          unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3494
        proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3495
          show "range \<sigma> \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3496
            using f_iter by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3497
          fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3498
          assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3499
          with \<open>k < 1\<close> \<open>f a \<noteq> a\<close> \<open>a \<in> M\<close> fim have gt0: "((1 - k) * \<epsilon>) / d a (f a) > 0"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  3500
            by (fastforce simp: divide_simps Pi_iff)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3501
          obtain N where "k^N < ((1-k) * \<epsilon>) / d a (f a)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3502
            using real_arch_pow_inv [OF gt0 \<open>k < 1\<close>] by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3503
          then have N: "\<And>n. n \<ge> N \<Longrightarrow> k^n < ((1-k) * \<epsilon>) / d a (f a)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3504
            by (smt (verit) \<open>0 < k\<close> assms(4) power_decreasing)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3505
          have "\<forall>n n'. n<n' \<longrightarrow> N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3506
          proof (intro exI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3507
            fix n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3508
            assume "n<n'" "N \<le> n" "N \<le> n'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3509
            have "d (\<sigma> n) (\<sigma> n') \<le> (\<Sum>i=n..<n'. d (\<sigma> i) (\<sigma> (Suc i)))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3510
            proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3511
              have "n < m \<Longrightarrow> d (\<sigma> n) (\<sigma> m) \<le> (\<Sum>i=n..<m. d (\<sigma> i) (\<sigma> (Suc i)))" for m
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3512
              proof (induction m)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3513
                case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3514
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3515
                  by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3516
              next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3517
                case (Suc m)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3518
                then consider "n<m" | "m=n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3519
                  by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3520
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3521
                proof cases
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3522
                  case 1
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3523
                  have "d (\<sigma> n) (\<sigma> (Suc m)) \<le> d (\<sigma> n) (\<sigma> m) + d (\<sigma> m) (\<sigma> (Suc m))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3524
                    by (simp add: f_iter triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3525
                  also have "\<dots> \<le> (\<Sum>i=n..<m. d (\<sigma> i) (\<sigma> (Suc i))) + d (\<sigma> m) (\<sigma> (Suc m))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3526
                    using Suc 1 by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3527
                  also have "\<dots> = (\<Sum>i = n..<Suc m. d (\<sigma> i) (\<sigma> (Suc i)))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3528
                    using "1" by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3529
                  finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3530
                qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3531
              qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3532
              with \<open>n < n'\<close> show ?thesis by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3533
            qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3534
            also have "\<dots> \<le> (\<Sum>i=n..<n'. d a (f a) * k^i)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3535
            proof (rule sum_mono)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3536
              fix i
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3537
              assume "i \<in> {n..<n'}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3538
              show "d (\<sigma> i) (\<sigma> (Suc i)) \<le> d a (f a) * k ^ i"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3539
              proof (induction i)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3540
                case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3541
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3542
                  by (auto simp: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3543
              next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3544
                case (Suc i)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3545
                have "d (\<sigma> (Suc i)) (\<sigma> (Suc (Suc i))) \<le> k * d (\<sigma> i) (\<sigma> (Suc i))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3546
                  using con \<sigma>_def f_iter fim by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3547
                also have "\<dots> \<le> d a (f a) * k ^ Suc i"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3548
                  using Suc \<open>0 < k\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3549
                finally show ?case .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3550
              qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3551
            qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3552
            also have "\<dots> = d a (f a) * (\<Sum>i=n..<n'. k^i)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3553
              by (simp add: sum_distrib_left)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3554
            also have "\<dots> = d a (f a) * (\<Sum>i=0..<n'-n. k^(i+n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3555
              using sum.shift_bounds_nat_ivl [of "power k" 0 n "n'-n"] \<open>n < n'\<close> by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3556
            also have "\<dots> = d a (f a) * k^n * (\<Sum>i<n'-n. k^i)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3557
              by (simp add: power_add lessThan_atLeast0 flip: sum_distrib_right)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3558
            also have "\<dots> = d a (f a) * (k ^ n - k ^ n') / (1 - k)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3559
              using \<open>k < 1\<close> \<open>n < n'\<close> apply (simp add: sum_gp_strict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3560
              by (simp add: algebra_simps flip: power_add)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3561
            also have "\<dots> < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3562
              using N \<open>k < 1\<close> \<open>0 < \<epsilon>\<close> \<open>0 < k\<close> \<open>N \<le> n\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3563
              apply (simp add: field_simps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3564
              by (smt (verit) nonneg pos_less_divide_eq zero_less_divide_iff zero_less_power)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3565
            finally show "d (\<sigma> n) (\<sigma> n') < \<epsilon>" .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3566
          qed 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3567
          then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (\<sigma> n) (\<sigma> n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3568
            by (metis \<open>0 < \<epsilon>\<close> commute f_iter linorder_not_le local.mdist_zero nat_less_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3569
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3570
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3571
      then obtain l where l: "limitin mtopology \<sigma> l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3572
        using \<open>mcomplete\<close> mcomplete_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3573
      show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3574
      proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3575
        show "l \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3576
          using l limitin_mspace by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3577
        show "f l = l"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3578
        proof (rule limitin_metric_unique)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3579
          have "limitin mtopology (f \<circ> \<sigma>) (f l) sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3580
          proof (rule continuous_map_limit)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3581
            have "Lipschitz_continuous_map Self Self f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3582
              using con by (auto simp: Lipschitz_continuous_map_def fim)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3583
            then show "continuous_map mtopology mtopology f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3584
              using Lipschitz_continuous_imp_continuous_map Self_def by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3585
          qed (use l in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3586
          moreover have "(f \<circ> \<sigma>) = (\<lambda>i. \<sigma>(i+1))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3587
            by (auto simp: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3588
          ultimately show "limitin mtopology (\<lambda>n. (f^^n)a) (f l) sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3589
            using limitin_sequentially_offset_rev [of mtopology \<sigma> 1]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3590
            by (simp add: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3591
        qed (use l in \<open>auto simp: \<sigma>_def\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3592
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3593
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3594
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3595
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3596
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3597
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3598
subsection\<open> The Baire Category Theorem\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3599
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3600
text \<open>Possibly relevant to the theorem "Baire" in Elementary Normed Spaces\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3601
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3602
lemma (in Metric_space) metric_Baire_category:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3603
  assumes "mcomplete" "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3604
  and "\<And>T. T \<in> \<G> \<Longrightarrow> openin mtopology T \<and> mtopology closure_of T = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3605
shows "mtopology closure_of \<Inter>\<G> = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3606
proof (cases "\<G>={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3607
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3608
  then obtain U :: "nat \<Rightarrow> 'a set" where U: "range U = \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3609
    by (metis \<open>countable \<G>\<close> uncountable_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3610
  with assms have u_open: "\<And>n. openin mtopology (U n)" and u_dense: "\<And>n. mtopology closure_of (U n) = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3611
    by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3612
  have "\<Inter>(range U) \<inter> W \<noteq> {}" if W: "openin mtopology W" "W \<noteq> {}" for W
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3613
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3614
    have "W \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3615
      using openin_mtopology W by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3616
    have "\<exists>r' x'. 0 < r' \<and> r' < r/2 \<and> x' \<in> M \<and> mcball x' r' \<subseteq> mball x r \<inter> U n" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3617
      if "r>0" "x \<in> M" for x r n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3618
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3619
      obtain z where z: "z \<in> U n \<inter> mball x r"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3620
        using u_dense [of n] \<open>r>0\<close> \<open>x \<in> M\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3621
        by (metis dense_intersects_open centre_in_mball_iff empty_iff openin_mball topspace_mtopology equals0I)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3622
      then have "z \<in> M" by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3623
      have "openin mtopology (U n \<inter> mball x r)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3624
        by (simp add: openin_Int u_open)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3625
      with \<open>z \<in> M\<close> z obtain e where "e>0" and e: "mcball z e \<subseteq> U n \<inter> mball x r"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3626
        by (meson openin_mtopology_mcball)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3627
      define r' where "r' \<equiv> min e (r/4)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3628
      show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3629
      proof (intro exI conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3630
        show "0 < r'" "r' < r / 2" "z \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3631
          using \<open>e>0\<close> \<open>r>0\<close> \<open>z \<in> M\<close> by (auto simp: r'_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3632
        show "mcball z r' \<subseteq> mball x r \<inter> U n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3633
          using Metric_space.mcball_subset_concentric e r'_def by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3634
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3635
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3636
    then obtain nextr nextx 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3637
      where nextr: "\<And>r x n. \<lbrakk>r>0; x\<in>M\<rbrakk> \<Longrightarrow> 0 < nextr r x n \<and> nextr r x n < r/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3638
        and nextx: "\<And>r x n. \<lbrakk>r>0; x\<in>M\<rbrakk> \<Longrightarrow> nextx r x n \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3639
        and nextsub: "\<And>r x n. \<lbrakk>r>0; x\<in>M\<rbrakk> \<Longrightarrow> mcball (nextx r x n) (nextr r x n) \<subseteq> mball x r \<inter> U n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3640
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3641
    obtain x0 where x0: "x0 \<in> U 0 \<inter> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3642
      by (metis W dense_intersects_open topspace_mtopology all_not_in_conv u_dense)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3643
    then have "x0 \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3644
      using \<open>W \<subseteq> M\<close> by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3645
    obtain r0 where "0 < r0" "r0 < 1" and sub: "mcball x0 r0 \<subseteq> U 0 \<inter> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3646
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3647
      have "openin mtopology (U 0 \<inter> W)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3648
        using W u_open by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3649
      then obtain r where "r>0" and r: "mball x0 r \<subseteq> U 0" "mball x0 r \<subseteq> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3650
        by (meson Int_subset_iff openin_mtopology x0)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3651
      define r0 where "r0 \<equiv> (min r 1) / 2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3652
      show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3653
      proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3654
        show "0 < r0" "r0 < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3655
          using \<open>r>0\<close> by (auto simp: r0_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3656
        show "mcball x0 r0 \<subseteq> U 0 \<inter> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3657
          using r \<open>0 < r0\<close> r0_def by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3658
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3659
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3660
    define b where "b \<equiv> rec_nat (x0,r0) (\<lambda>n (x,r). (nextx r x n, nextr r x n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3661
    have b0[simp]: "b 0 = (x0,r0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3662
      by (simp add: b_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3663
    have bSuc[simp]: "b (Suc n) = (let (x,r) = b n in (nextx r x n, nextr r x n))" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3664
      by (simp add: b_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3665
    define xf where "xf \<equiv> fst \<circ> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3666
    define rf where "rf \<equiv> snd \<circ> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3667
    have rfxf: "0 < rf n \<and> xf n \<in> M" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3668
    proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3669
      case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3670
      with \<open>0 < r0\<close> \<open>x0 \<in> M\<close> show ?case 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3671
        by (auto simp: rf_def xf_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3672
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3673
      case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3674
      then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3675
        by (auto simp: rf_def xf_def case_prod_unfold nextr nextx Let_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3676
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3677
    have mcball_sub: "mcball (xf (Suc n)) (rf (Suc n)) \<subseteq> mball (xf n) (rf n) \<inter> U n" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3678
      using rfxf nextsub by (auto simp: xf_def rf_def case_prod_unfold Let_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3679
    have half: "rf (Suc n) < rf n / 2" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3680
      using rfxf nextr by (auto simp: xf_def rf_def case_prod_unfold Let_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3681
    then have "decseq rf"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3682
      using rfxf by (smt (verit, ccfv_threshold) decseq_SucI field_sum_of_halves)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3683
    have nested: "mball (xf n) (rf n) \<subseteq> mball (xf m) (rf m)" if "m \<le> n" for m n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3684
      using that
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3685
    proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3686
      case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3687
      then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3688
        by (metis mcball_sub order.trans inf.boundedE le_Suc_eq mball_subset_mcball order.refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3689
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3690
    have "MCauchy xf"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3691
      unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3692
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3693
      show "range xf \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3694
        using rfxf by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3695
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3696
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3697
      then obtain N where N: "inverse (2^N) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3698
        using real_arch_pow_inv by (force simp flip: power_inverse)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3699
      have "d (xf n) (xf n') < \<epsilon>" if "n \<le> n'" "N \<le> n" "N \<le> n'" for n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3700
      proof -           
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3701
        have *: "rf n < inverse (2 ^ n)" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3702
        proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3703
          case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3704
          then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3705
            by (simp add: \<open>r0 < 1\<close> rf_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3706
        next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3707
          case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3708
          with half show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3709
            by simp (smt (verit))
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3710
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3711
        have "rf n \<le> rf N"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3712
          using \<open>decseq rf\<close> \<open>N \<le> n\<close> by (simp add: decseqD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3713
        moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3714
        have "xf n' \<in> mball (xf n) (rf n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3715
          using nested rfxf \<open>n \<le> n'\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3716
        ultimately have "d (xf n) (xf n') < rf N"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3717
          by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3718
        also have "\<dots> < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3719
          using "*" N order.strict_trans by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3720
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3721
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3722
      then show "\<exists>N. \<forall>n n'. N \<le> n \<longrightarrow> N \<le> n' \<longrightarrow> d (xf n) (xf n') < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3723
        by (metis commute linorder_le_cases)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3724
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3725
    then obtain l where l: "limitin mtopology xf l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3726
      using \<open>mcomplete\<close> mcomplete_alt by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3727
    have l_in: "l \<in> mcball (xf n) (rf n)" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3728
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3729
      have "\<forall>\<^sub>F m in sequentially. xf m \<in> mcball (xf n) (rf n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3730
        unfolding eventually_sequentially
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3731
        by (meson nested rfxf centre_in_mball_iff mball_subset_mcball subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3732
      with l limitin_closedin show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3733
        by (metis closedin_mcball trivial_limit_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3734
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3735
    then have "\<And>n. l \<in> U n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3736
      using mcball_sub by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3737
    moreover have "l \<in> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3738
      using l_in[of 0] sub  by (auto simp: xf_def rf_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3739
    ultimately show ?thesis by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3740
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3741
  with U show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3742
    by (metis dense_intersects_open topspace_mtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3743
qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3744
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3745
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3746
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3747
lemma (in Metric_space) metric_Baire_category_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3748
  assumes "mcomplete" "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3749
    and empty: "\<And>T. T \<in> \<G> \<Longrightarrow> closedin mtopology T \<and> mtopology interior_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3750
  shows "mtopology interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3751
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3752
  have *: "mtopology closure_of \<Inter>((-)M ` \<G>) = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3753
  proof (intro metric_Baire_category conjI \<open>mcomplete\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3754
    show "countable ((-) M ` \<G>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3755
      using \<open>countable \<G>\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3756
    fix T
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3757
    assume "T \<in> (-) M ` \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3758
    then obtain U where U: "U \<in> \<G>" "T = M-U" "U \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3759
      using empty metric_closedin_iff_sequentially_closed by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3760
    with empty show "openin mtopology T" by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3761
    show "mtopology closure_of T = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3762
      using U by (simp add: closure_of_interior_of double_diff empty)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3763
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3764
  with closure_of_eq show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3765
    by (fastforce simp: interior_of_closure_of split: if_split_asm)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3766
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3767
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3768
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3769
text \<open>Since all locally compact Hausdorff spaces are regular, the disjunction in the HOL Light version is redundant.\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3770
lemma Baire_category_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3771
  assumes "locally_compact_space X" "regular_space X" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3772
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3773
    and empty: "\<And>G. G \<in> \<G> \<Longrightarrow> closedin X G \<and> X interior_of G = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3774
  shows "X interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3775
proof (cases "\<G> = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3776
  case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3777
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3778
    by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3779
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3780
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3781
  then obtain T :: "nat \<Rightarrow> 'a set" where T: "\<G> = range T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3782
    by (metis \<open>countable \<G>\<close> uncountable_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3783
  with empty have Tempty: "\<And>n. X interior_of (T n) = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3784
    by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3785
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3786
  proof (clarsimp simp: T interior_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3787
    fix z U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3788
    assume "z \<in> U" and opeA: "openin X U" and Asub: "U \<subseteq> \<Union> (range T)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3789
    with openin_subset have "z \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3790
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3791
    have "neighbourhood_base_of (\<lambda>C. compactin X C \<and> closedin X C) X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3792
      using assms locally_compact_regular_space_neighbourhood_base by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3793
    then obtain V K where "openin X V" "compactin X K" "closedin X K" "z \<in> V" "V \<subseteq> K" "K \<subseteq> U"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3794
      by (metis (no_types, lifting) \<open>z \<in> U\<close> neighbourhood_base_of opeA)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3795
    have nb_closedin: "neighbourhood_base_of (closedin X) X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3796
      using \<open>regular_space X\<close> neighbourhood_base_of_closedin by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3797
    have "\<exists>\<Phi>. \<forall>n. (\<Phi> n \<subseteq> K \<and> closedin X (\<Phi> n) \<and> X interior_of \<Phi> n \<noteq> {} \<and> disjnt (\<Phi> n) (T n)) \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3798
        \<Phi> (Suc n) \<subseteq> \<Phi> n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3799
    proof (rule dependent_nat_choice)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3800
      show "\<exists>x\<subseteq>K. closedin X x \<and> X interior_of x \<noteq> {} \<and> disjnt x (T 0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3801
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3802
        have False if "V \<subseteq> T 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3803
          using Tempty \<open>openin X V\<close> \<open>z \<in> V\<close> interior_of_maximal that by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3804
        then obtain x where "openin X (V - T 0) \<and> x \<in> V - T 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3805
          using T \<open>openin X V\<close> empty by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3806
        with nb_closedin 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3807
        obtain N C  where "openin X N" "closedin X C" "x \<in> N" "N \<subseteq> C" "C \<subseteq> V - T 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3808
          unfolding neighbourhood_base_of by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3809
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3810
        proof (intro exI conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3811
          show "C \<subseteq> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3812
            using \<open>C \<subseteq> V - T 0\<close> \<open>V \<subseteq> K\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3813
          show "X interior_of C \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3814
            by (metis \<open>N \<subseteq> C\<close> \<open>openin X N\<close> \<open>x \<in> N\<close> empty_iff interior_of_eq_empty)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3815
          show "disjnt C (T 0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3816
            using \<open>C \<subseteq> V - T 0\<close> disjnt_iff by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3817
        qed (use \<open>closedin X C\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3818
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3819
      show "\<exists>L. (L \<subseteq> K \<and> closedin X L \<and> X interior_of L \<noteq> {} \<and> disjnt L (T (Suc n))) \<and> L \<subseteq> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3820
        if \<section>: "C \<subseteq> K \<and> closedin X C \<and> X interior_of C \<noteq> {} \<and> disjnt C (T n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3821
        for C n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3822
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3823
        have False if "X interior_of C \<subseteq> T (Suc n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3824
          by (metis Tempty interior_of_eq_empty \<section> openin_interior_of that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3825
        then obtain x where "openin X (X interior_of C - T (Suc n)) \<and> x \<in> X interior_of C - T (Suc n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3826
          using T empty by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3827
        with nb_closedin 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3828
        obtain N D where "openin X N" "closedin X D" "x \<in> N" "N \<subseteq> D"  and D: "D \<subseteq> X interior_of C - T(Suc n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3829
          unfolding neighbourhood_base_of by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3830
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3831
        proof (intro conjI exI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3832
          show "D \<subseteq> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3833
            using D interior_of_subset \<section> by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3834
          show "X interior_of D \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3835
            by (metis \<open>N \<subseteq> D\<close> \<open>openin X N\<close> \<open>x \<in> N\<close> empty_iff interior_of_eq_empty)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3836
          show "disjnt D (T (Suc n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3837
            using D disjnt_iff by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3838
          show "D \<subseteq> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3839
            using interior_of_subset [of X C] D by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3840
        qed (use \<open>closedin X D\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3841
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3842
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3843
    then obtain \<Phi> where \<Phi>: "\<And>n. \<Phi> n \<subseteq> K \<and> closedin X (\<Phi> n) \<and> X interior_of \<Phi> n \<noteq> {} \<and> disjnt (\<Phi> n) (T n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3844
      and "\<And>n. \<Phi> (Suc n) \<subseteq> \<Phi> n" by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3845
    then have "decseq \<Phi>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3846
      by (simp add: decseq_SucI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3847
    moreover have "\<And>n. \<Phi> n \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3848
      by (metis \<Phi> bot.extremum_uniqueI interior_of_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3849
    ultimately have "\<Inter>(range \<Phi>) \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3850
      by (metis \<Phi> compact_space_imp_nest \<open>compactin X K\<close> compactin_subspace closedin_subset_topspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3851
    moreover have "U \<subseteq> {y. \<exists>x. y \<in> T x}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3852
      using Asub by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3853
    with T have "{a. \<forall>n.  a \<in> \<Phi> n} \<subseteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3854
      by (smt (verit) Asub \<Phi> Collect_empty_eq UN_iff \<open>K \<subseteq> U\<close> disjnt_iff subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3855
    ultimately show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3856
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3857
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3858
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3859
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3860
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3861
lemma Baire_category_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3862
  assumes "completely_metrizable_space X \<or> locally_compact_space X \<and> regular_space X" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3863
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3864
    and "\<And>T. T \<in> \<G> \<Longrightarrow> closedin X T \<and> X interior_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3865
  shows "X interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3866
  using Baire_category_aux [of X \<G>] Metric_space.metric_Baire_category_alt
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3867
  by (metis assms completely_metrizable_space_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3868
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3869
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3870
lemma Baire_category:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3871
  assumes "completely_metrizable_space X \<or> locally_compact_space X \<and> regular_space X" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3872
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3873
    and top: "\<And>T. T \<in> \<G> \<Longrightarrow> openin X T \<and> X closure_of T = topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3874
  shows "X closure_of \<Inter>\<G> = topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3875
proof (cases "\<G>={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3876
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3877
  have *: "X interior_of \<Union>((-)(topspace X) ` \<G>) = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3878
    proof (intro Baire_category_alt conjI assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3879
  show "countable ((-) (topspace X) ` \<G>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3880
    using assms by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3881
    fix T
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3882
    assume "T \<in> (-) (topspace X) ` \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3883
    then obtain U where U: "U \<in> \<G>" "T = (topspace X) - U" "U \<subseteq> (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3884
      by (meson top image_iff openin_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3885
    then show "closedin X T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3886
      by (simp add: closedin_diff top)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3887
    show "X interior_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3888
      using U top by (simp add: interior_of_closure_of double_diff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3889
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3890
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3891
    by (simp add: closure_of_eq_topspace interior_of_complement)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3892
qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3893
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3894
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3895
subsection\<open>Sierpinski-Hausdorff type results about countable closed unions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3896
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3897
lemma locally_connected_not_countable_closed_union:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3898
  assumes "topspace X \<noteq> {}" and csX: "connected_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3899
    and lcX: "locally_connected_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3900
    and X: "completely_metrizable_space X \<or> locally_compact_space X \<and> Hausdorff_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3901
    and "countable \<U>" and pwU: "pairwise disjnt \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3902
    and clo: "\<And>C. C \<in> \<U> \<Longrightarrow> closedin X C \<and> C \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3903
    and UU_eq: "\<Union>\<U> = topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3904
  shows "\<U> = {topspace X}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3905
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3906
  define \<V> where "\<V> \<equiv> (frontier_of) X ` \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3907
  define B where "B \<equiv> \<Union>\<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3908
  then have Bsub: "B \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3909
    by (simp add: Sup_le_iff \<V>_def closedin_frontier_of closedin_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3910
  have allsub: "A \<subseteq> topspace X" if "A \<in> \<U>" for A
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3911
    by (meson clo closedin_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3912
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3913
  proof (rule ccontr)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3914
    assume "\<U> \<noteq> {topspace X}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3915
    with assms have "\<exists>A\<in>\<U>. \<not> (closedin X A \<and> openin X A)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3916
      by (metis Union_empty connected_space_clopen_in singletonI subsetI subset_singleton_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3917
    then have "B \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3918
      by (auto simp: B_def \<V>_def frontier_of_eq_empty allsub)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3919
    moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3920
    have "subtopology X B interior_of B = B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3921
      by (simp add: Bsub interior_of_openin openin_subtopology_refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3922
    ultimately have int_B_nonempty: "subtopology X B interior_of B \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3923
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3924
    have "subtopology X B interior_of \<Union>\<V> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3925
    proof (intro Baire_category_alt conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3926
      have "\<Union>\<U> \<subseteq> B \<union> \<Union>((interior_of) X ` \<U>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3927
        using clo closure_of_closedin by (fastforce simp: B_def \<V>_def frontier_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3928
      moreover have "B \<union> \<Union>((interior_of) X ` \<U>) \<subseteq> \<Union>\<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3929
        using allsub clo frontier_of_subset_eq interior_of_subset by (fastforce simp: B_def \<V>_def )
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3930
      moreover have "disjnt B (\<Union>((interior_of) X ` \<U>))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3931
        using pwU 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3932
        apply (clarsimp simp: B_def \<V>_def frontier_of_def pairwise_def disjnt_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3933
        by (metis clo closure_of_eq interior_of_subset subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3934
      ultimately have "B = topspace X - \<Union> ((interior_of) X ` \<U>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3935
        by (auto simp: UU_eq disjnt_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3936
      then have "closedin X B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3937
        by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3938
      with X show "completely_metrizable_space (subtopology X B) \<or> locally_compact_space (subtopology X B) \<and> regular_space (subtopology X B)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3939
        by (metis completely_metrizable_space_closedin locally_compact_Hausdorff_or_regular 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3940
            locally_compact_space_closed_subset regular_space_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3941
      show "countable \<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3942
        by (simp add: \<V>_def \<open>countable \<U>\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3943
      fix V
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3944
      assume "V \<in> \<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3945
      then obtain S where S: "S \<in> \<U>" "V = X frontier_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3946
        by (auto simp: \<V>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3947
      show "closedin (subtopology X B) V"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3948
        by (metis B_def Sup_upper \<V>_def \<open>V \<in> \<V>\<close> closedin_frontier_of closedin_subset_topspace image_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3949
      have "subtopology X B interior_of (X frontier_of S) = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3950
      proof (clarsimp simp: interior_of_def openin_subtopology_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3951
        fix a U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3952
        assume "a \<in> B" "a \<in> U" and opeU: "openin X U" and BUsub: "B \<inter> U \<subseteq> X frontier_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3953
        then have "a \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3954
          by (meson IntI \<open>S \<in> \<U>\<close> clo frontier_of_subset_closedin subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3955
        then obtain W C where "openin X W" "connectedin X C" "a \<in> W" "W \<subseteq> C" "C \<subseteq> U"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3956
          by (metis \<open>a \<in> U\<close> lcX locally_connected_space opeU)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3957
        have "W \<inter> X frontier_of S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3958
          using \<open>B \<inter> U \<subseteq> X frontier_of S\<close> \<open>a \<in> B\<close> \<open>a \<in> U\<close> \<open>a \<in> W\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3959
        with frontier_of_openin_straddle_Int
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3960
        obtain "W \<inter> S \<noteq> {}" "W - S \<noteq> {}" "W \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3961
          using \<open>openin X W\<close> by (metis openin_subset) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3962
        then obtain b where "b \<in> topspace X" "b \<in> W-S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3963
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3964
        with UU_eq obtain T where "T \<in> \<U>" "T \<noteq> S" "W \<inter> T \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3965
          by auto 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3966
        then have "disjnt S T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3967
          by (metis \<open>S \<in> \<U>\<close> pairwise_def pwU)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3968
        then have "C - T \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3969
          by (meson Diff_eq_empty_iff \<open>W \<subseteq> C\<close> \<open>a \<in> S\<close> \<open>a \<in> W\<close> disjnt_iff subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3970
        then have "C \<inter> X frontier_of T \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3971
          using \<open>W \<inter> T \<noteq> {}\<close> \<open>W \<subseteq> C\<close> \<open>connectedin X C\<close> connectedin_Int_frontier_of by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3972
        moreover have "C \<inter> X frontier_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3973
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3974
          have "X frontier_of S \<subseteq> S" "X frontier_of T \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3975
            using frontier_of_subset_closedin \<open>S \<in> \<U>\<close> \<open>T \<in> \<U>\<close> clo by blast+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3976
          moreover have "X frontier_of T \<union> B = B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3977
            using B_def \<V>_def \<open>T \<in> \<U>\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3978
          ultimately show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3979
            using BUsub \<open>C \<subseteq> U\<close> \<open>disjnt S T\<close> unfolding disjnt_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3980
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3981
        ultimately show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3982
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3983
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3984
      with S show "subtopology X B interior_of V = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3985
        by meson
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3986
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3987
    then show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3988
      using B_def int_B_nonempty by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3989
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3990
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3991
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3992
lemma real_Sierpinski_lemma:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3993
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3994
  assumes "a \<le> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3995
    and "countable \<U>" and pwU: "pairwise disjnt \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3996
    and clo: "\<And>C. C \<in> \<U> \<Longrightarrow> closed C \<and> C \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3997
    and "\<Union>\<U> = {a..b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3998
  shows "\<U> = {{a..b}}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3999
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4000
  have "locally_connected_space (top_of_set {a..b})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4001
    by (simp add: locally_connected_real_interval)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4002
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4003
  have "completely_metrizable_space (top_of_set {a..b})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4004
    by (metis box_real(2) completely_metrizable_space_cbox)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4005
  ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4006
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4007
    using locally_connected_not_countable_closed_union [of "subtopology euclidean {a..b}"] assms
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4008
    apply (simp add: closedin_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4009
    by (metis Union_upper inf.orderE)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4010
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4011
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4012
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4013
subsection\<open>The Tychonoff embedding\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4014
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4015
lemma completely_regular_space_cube_embedding_explicit:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4016
  assumes "completely_regular_space X" "Hausdorff_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4017
  shows "embedding_map X
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4018
             (product_topology (\<lambda>f. top_of_set {0..1::real})
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4019
                (mspace (submetric (cfunspace X euclidean_metric)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4020
                  {f. f \<in> topspace X \<rightarrow> {0..1}})) )
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4021
             (\<lambda>x. \<lambda>f \<in> mspace (submetric (cfunspace X euclidean_metric) {f. f \<in> topspace X \<rightarrow> {0..1}}).
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4022
              f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4023
proof -
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4024
  define K where "K \<equiv> mspace(submetric (cfunspace X euclidean_metric) {f. f \<in> topspace X \<rightarrow> {0..1::real}})"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4025
  define e where "e \<equiv> \<lambda>x. \<lambda>f\<in>K. f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4026
  have "e x \<noteq> e y" if xy: "x \<noteq> y" "x \<in> topspace X" "y \<in> topspace X" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4027
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4028
    have "closedin X {x}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4029
      by (simp add: \<open>Hausdorff_space X\<close> closedin_Hausdorff_singleton \<open>x \<in> topspace X\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4030
    then obtain f where contf: "continuous_map X euclideanreal f" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4031
      and f01: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> {0..1}" and fxy: "f y = 0" "f x = 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4032
      using \<open>completely_regular_space X\<close> xy unfolding completely_regular_space_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4033
      by (smt (verit, ccfv_threshold) Diff_iff continuous_map_in_subtopology image_subset_iff singleton_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4034
    then have "bounded (f ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4035
      by (meson bounded_closed_interval bounded_subset image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4036
    with contf f01 have "restrict f (topspace X) \<in> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4037
      by (auto simp: K_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4038
    with fxy xy show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4039
      unfolding e_def by (metis restrict_apply' zero_neq_one)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4040
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4041
  then have "inj_on e (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4042
    by (meson inj_onI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4043
  then obtain e' where e': "\<And>x. x \<in> topspace X \<Longrightarrow> e' (e x) = x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4044
    by (metis inv_into_f_f)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4045
  have "continuous_map (subtopology (product_topology (\<lambda>f. top_of_set {0..1}) K) (e ` topspace X)) X e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4046
  proof (clarsimp simp add: continuous_map_atin limitin_atin openin_subtopology_alt e')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4047
    fix x U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4048
    assume "e x \<in> K \<rightarrow>\<^sub>E {0..1}" and "x \<in> topspace X" and "openin X U" and "x \<in> U"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4049
    then obtain g where contg: "continuous_map X (top_of_set {0..1}) g" and "g x = 0" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4050
          and gim: "g ` (topspace X - U) \<subseteq> {1::real}"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4051
      using \<open>completely_regular_space X\<close> unfolding completely_regular_space_def 
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4052
      by (metis Diff_iff openin_closedin_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4053
    then have "bounded (g ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4054
      by (meson bounded_closed_interval bounded_subset continuous_map_in_subtopology)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4055
    moreover have "g \<in> topspace X \<rightarrow> {0..1}"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4056
      using contg by (simp add: continuous_map_def)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4057
    ultimately have g_in_K: "restrict g (topspace X) \<in> K"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4058
      using contg by (force simp add: K_def continuous_map_in_subtopology)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4059
    have "openin (top_of_set {0..1}) {0..<1::real}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4060
      using open_real_greaterThanLessThan[of "-1" 1] by (force simp: openin_open)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4061
    moreover have "e x \<in> (\<Pi>\<^sub>E f\<in>K. if f = restrict g (topspace X) then {0..<1} else {0..1})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4062
      using \<open>e x \<in> K \<rightarrow>\<^sub>E {0..1}\<close> by (simp add: e_def \<open>g x = 0\<close> \<open>x \<in> topspace X\<close> PiE_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4063
    moreover have "e y = e x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4064
      if "y \<notin> U" and ey: "e y \<in> (\<Pi>\<^sub>E f\<in>K. if f = restrict g (topspace X) then {0..<1} else {0..1})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4065
           and y: "y \<in> topspace X" for y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4066
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4067
      have "e y (restrict g (topspace X)) \<in> {0..<1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4068
        using ey by (smt (verit, ccfv_SIG) PiE_mem g_in_K)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4069
    with gim g_in_K y \<open>y \<notin> U\<close> show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4070
      by (fastforce simp: e_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4071
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4072
    ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4073
    show "\<exists>W. openin (product_topology (\<lambda>f. top_of_set {0..1}) K) W \<and> e x \<in> W \<and> e' ` (e ` topspace X \<inter> W - {e x}) \<subseteq> U"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4074
      apply (rule_tac x="PiE K (\<lambda>f. if f = restrict g (topspace X) then {0..<1} else {0..1})" in exI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4075
      by (auto simp: openin_PiE_gen e')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4076
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4077
  with e' have "embedding_map X (product_topology (\<lambda>f. top_of_set {0..1}) K) e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4078
    unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4079
    by (fastforce simp: e_def K_def continuous_map_in_subtopology continuous_map_componentwise)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4080
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4081
    by (simp add: K_def e_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4082
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4083
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4084
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4085
lemma completely_regular_space_cube_embedding:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4086
  fixes X :: "'a topology"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4087
  assumes "completely_regular_space X" "Hausdorff_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4088
  obtains K:: "('a\<Rightarrow>real)set" and e
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4089
    where "embedding_map X (product_topology (\<lambda>f. top_of_set {0..1::real}) K) e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4090
  using completely_regular_space_cube_embedding_explicit [OF assms] by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4091
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4092
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4093
subsection \<open>Urysohn and Tietze analogs for completely regular spaces\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4094
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4095
text\<open>"Urysohn and Tietze analogs for completely regular spaces if (()) set is 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4096
assumed compact instead of closed. Note that Hausdorffness is *not* required: inside () proof 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4097
we factor through the Kolmogorov quotient." -- John Harrison\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4098
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4099
lemma Urysohn_completely_regular_closed_compact:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4100
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4101
  assumes "a \<le> b" "completely_regular_space X" "closedin X S" "compactin X T" "disjnt S T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4102
  obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T \<subseteq> {a}" "f ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4103
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4104
  obtain f where contf: "continuous_map X (subtopology euclideanreal {0..1}) f" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4105
    and f0: "f ` T \<subseteq> {0}" and f1: "f ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4106
  proof (cases "T={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4107
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4108
    show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4109
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4110
      show "continuous_map X (top_of_set {0..1}) (\<lambda>x. 1::real)" "(\<lambda>x. 1::real) ` T \<subseteq> {0}" "(\<lambda>x. 1::real) ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4111
        using True by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4112
    qed   
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4113
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4114
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4115
    have "\<And>t. t \<in> T \<Longrightarrow> \<exists>f. continuous_map X (subtopology euclideanreal ({0..1})) f \<and> f t = 0 \<and> f ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4116
      using assms unfolding completely_regular_space_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4117
      by (meson DiffI compactin_subset_topspace disjnt_iff subset_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4118
    then obtain g where contg: "\<And>t. t \<in> T \<Longrightarrow> continuous_map X (subtopology euclideanreal {0..1}) (g t)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4119
                  and g0: "\<And>t. t \<in> T \<Longrightarrow> g t t = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4120
                  and g1: "\<And>t. t \<in> T \<Longrightarrow> g t ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4121
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4122
    then have g01: "\<And>t. t \<in> T \<Longrightarrow> g t ` topspace X \<subseteq> {0..1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4123
      by (meson continuous_map_in_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4124
    define G where "G \<equiv> \<lambda>t. {x \<in> topspace X. g t x \<in> {..<1/2}}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4125
    have "Ball (G`T) (openin X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4126
      using contg unfolding G_def continuous_map_in_subtopology
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4127
      by (smt (verit, best) Collect_cong openin_continuous_map_preimage image_iff open_lessThan open_openin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4128
    moreover have "T \<subseteq> \<Union>(G`T)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4129
      using \<open>compactin X T\<close> g0 compactin_subset_topspace by (force simp: G_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4130
    ultimately have "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> G`T \<and> T \<subseteq> \<Union> \<F>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4131
      using \<open>compactin X T\<close> unfolding compactin_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4132
    then obtain K where K: "finite K" "K \<subseteq> T" "T \<subseteq> \<Union>(G`K)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4133
      by (metis finite_subset_image)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4134
    with False have "K \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4135
      by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4136
    define f where "f \<equiv> \<lambda>x. 2 * max 0 (Inf ((\<lambda>t. g t x) ` K) - 1/2)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4137
    have [simp]: "max 0 (x - 1/2) = 0 \<longleftrightarrow> x \<le> 1/2" for x::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4138
      by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4139
    have [simp]: "2 * max 0 (x - 1/2) = 1 \<longleftrightarrow> x = 1" for x::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4140
      by (simp add: max_def_raw)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4141
    show thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4142
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4143
      have "g t s = 1" if "s \<in> S" "t \<in> K" for s t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4144
        using \<open>K \<subseteq> T\<close> g1 that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4145
      then show "f ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4146
        using \<open>K \<noteq> {}\<close> by (simp add: f_def image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4147
      have "(INF t\<in>K. g t x) \<le> 1/2" if "x \<in> T" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4148
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4149
        obtain k where "k \<in> K" "g k x < 1/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4150
          using K \<open>x \<in> T\<close> by (auto simp: G_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4151
        then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4152
          by (meson \<open>finite K\<close> cInf_le_finite dual_order.trans finite_imageI imageI less_le_not_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4153
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4154
      then show "f ` T \<subseteq> {0}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4155
        by (force simp: f_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4156
      have "\<And>t. t \<in> K \<Longrightarrow> continuous_map X euclideanreal (g t)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4157
        using \<open>K \<subseteq> T\<close> contg continuous_map_in_subtopology by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4158
      moreover have "2 * max 0 ((INF t\<in>K. g t x) - 1/2) \<le> 1" if "x \<in> topspace X" for x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4159
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4160
        obtain k where "k \<in> K" "g k x \<le> 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4161
          using K \<open>x \<in> topspace X\<close> \<open>K \<noteq> {}\<close> g01 by (fastforce simp: G_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4162
        then have "(INF t\<in>K. g t x) \<le> 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4163
          by (meson \<open>finite K\<close> cInf_le_finite dual_order.trans finite_imageI imageI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4164
        then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4165
          by (simp add: max_def_raw)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4166
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4167
      ultimately show "continuous_map X (top_of_set {0..1}) f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4168
        by (force simp: f_def continuous_map_in_subtopology intro!: \<open>finite K\<close> continuous_intros)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4169
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4170
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4171
  define g where "g \<equiv> \<lambda>x. a + (b - a) * f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4172
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4173
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4174
    have "\<forall>x\<in>topspace X. a + (b - a) * f x \<le> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4175
      using contf \<open>a \<le> b\<close> apply (simp add: continuous_map_in_subtopology image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4176
      by (smt (verit, best) mult_right_le_one_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4177
    then show "continuous_map X (top_of_set {a..b}) g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4178
      using contf \<open>a \<le> b\<close> unfolding g_def continuous_map_in_subtopology image_subset_iff
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4179
      by (intro conjI continuous_intros; simp)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4180
    show "g ` T \<subseteq> {a}" "g ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4181
      using f0 f1 by (auto simp: g_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4182
  qed 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4183
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4184
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4185
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4186
lemma Urysohn_completely_regular_compact_closed:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4187
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4188
  assumes "a \<le> b" "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4189
  obtains f where "continuous_map X (subtopology euclidean {a..b}) f" "f ` T \<subseteq> {a}" "f ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4190
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4191
  obtain f where contf: "continuous_map X (subtopology euclidean {-b..-a}) f" and fim: "f ` T \<subseteq> {-a}" "f ` S \<subseteq> {-b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4192
    by (meson Urysohn_completely_regular_closed_compact assms disjnt_sym neg_le_iff_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4193
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4194
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4195
    show "continuous_map X (top_of_set {a..b}) (uminus \<circ> f)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4196
      using contf by (auto simp: continuous_map_in_subtopology o_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4197
    show "(uminus o f) ` T \<subseteq> {a}" "(uminus o f) ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4198
      using fim by fastforce+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4199
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4200
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4201
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4202
lemma Urysohn_completely_regular_compact_closed_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4203
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4204
  assumes "completely_regular_space X" "compactin X S" "closedin X T" "disjnt S T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4205
  obtains f where "continuous_map X euclideanreal f" "f ` T \<subseteq> {a}" "f ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4206
proof (cases a b rule: le_cases)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4207
  case le
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4208
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4209
    by (meson Urysohn_completely_regular_compact_closed assms continuous_map_into_fulltopology that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4210
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4211
  case ge
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4212
  then show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4213
    using Urysohn_completely_regular_closed_compact assms
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4214
    by (metis Urysohn_completely_regular_closed_compact assms continuous_map_into_fulltopology disjnt_sym that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4215
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4216
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4217
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4218
lemma Tietze_extension_comp_reg_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4219
  fixes T :: "real set"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4220
  assumes "completely_regular_space X" "Hausdorff_space X" "compactin X S" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4221
    and T: "is_interval T" "T\<noteq>{}" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4222
    and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4223
  obtains g where "continuous_map X euclidean g" "g ` topspace X \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4224
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4225
  obtain K:: "('a\<Rightarrow>real)set" and e
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4226
    where e0: "embedding_map X (product_topology (\<lambda>f. top_of_set {0..1::real}) K) e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4227
    using assms completely_regular_space_cube_embedding by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4228
  define cube where "cube \<equiv> product_topology (\<lambda>f. top_of_set {0..1::real}) K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4229
  have e: "embedding_map X cube e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4230
    using e0 by (simp add: cube_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4231
  obtain e' where  e': "homeomorphic_maps X (subtopology cube (e ` topspace X)) e e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4232
    using e by (force simp: cube_def embedding_map_def homeomorphic_map_maps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4233
  then have conte: "continuous_map X (subtopology cube (e ` topspace X)) e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4234
     and conte': "continuous_map (subtopology cube (e ` topspace X)) X e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4235
     and e'e: "\<forall>x\<in>topspace X. e' (e x) = x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4236
    by (auto simp: homeomorphic_maps_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4237
  have "Hausdorff_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4238
    unfolding cube_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4239
    using Hausdorff_space_euclidean Hausdorff_space_product_topology Hausdorff_space_subtopology by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4240
  have "normal_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4241
  proof (rule compact_Hausdorff_or_regular_imp_normal_space)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4242
    show "compact_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4243
      unfolding cube_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4244
      using compact_space_product_topology compact_space_subtopology compactin_euclidean_iff by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4245
  qed (use \<open>Hausdorff_space cube\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4246
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4247
  have comp: "compactin cube (e ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4248
    by (meson \<open>compactin X S\<close> conte continuous_map_in_subtopology image_compactin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4249
  then have "closedin cube (e ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4250
    by (intro compactin_imp_closedin \<open>Hausdorff_space cube\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4251
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4252
  have "continuous_map (subtopology cube (e ` S)) euclideanreal (f \<circ> e')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4253
  proof (intro continuous_map_compose)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4254
    show "continuous_map (subtopology cube (e ` S)) (subtopology X S) e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4255
      unfolding continuous_map_in_subtopology
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4256
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4257
      show "continuous_map (subtopology cube (e ` S)) X e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4258
        by (meson \<open>compactin X S\<close> compactin_subset_topspace conte' continuous_map_from_subtopology_mono image_mono)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4259
      show "e' ` topspace (subtopology cube (e ` S)) \<subseteq> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4260
        using \<open>compactin X S\<close> compactin_subset_topspace e'e by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4261
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4262
  qed (simp add: contf)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4263
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4264
  have "(f \<circ> e') ` e ` S \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4265
    using \<open>compactin X S\<close> compactin_subset_topspace e'e fim by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4266
  ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4267
  obtain g where contg: "continuous_map cube euclidean g" and gsub: "g ` topspace cube \<subseteq> T" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4268
                and gf: "\<And>x. x \<in> e`S \<Longrightarrow> g x = (f \<circ> e') x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4269
    using Tietze_extension_realinterval T by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4270
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4271
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4272
    show "continuous_map X euclideanreal (g \<circ> e)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4273
      by (meson contg conte continuous_map_compose continuous_map_in_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4274
    show "(g \<circ> e) ` topspace X \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4275
      using gsub conte continuous_map_image_subset_topspace by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4276
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4277
    assume "x \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4278
    then show "(g \<circ> e) x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4279
      using gf \<open>compactin X S\<close> compactin_subset_topspace e'e by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4280
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4281
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4282
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4283
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4284
lemma Tietze_extension_completely_regular:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4285
  assumes "completely_regular_space X" "compactin X S" "is_interval T" "T \<noteq> {}" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4286
    and contf: "continuous_map (subtopology X S) euclidean f" and fim: "f`S \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4287
  obtains g where "continuous_map X euclideanreal g" "g ` topspace X \<subseteq> T" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4288
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4289
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4290
  define Q where "Q \<equiv> Kolmogorov_quotient X ` (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4291
  obtain g where contg: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) euclidean g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4292
    and gf: "\<And>x. x \<in> S \<Longrightarrow> g(Kolmogorov_quotient X x) = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4293
    using Kolmogorov_quotient_lift_exists 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4294
    by (metis \<open>compactin X S\<close> contf compactin_subset_topspace open_openin t0_space_def t1_space)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4295
  have "S \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4296
    by (simp add: \<open>compactin X S\<close> compactin_subset_topspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4297
  then have [simp]: "Q \<inter> Kolmogorov_quotient X ` S = Kolmogorov_quotient X ` S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4298
    using Q_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4299
  have creg: "completely_regular_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4300
    by (simp add: \<open>completely_regular_space X\<close> completely_regular_space_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4301
  then have "regular_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4302
    by (simp add: completely_regular_imp_regular_space)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4303
  then have "Hausdorff_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4304
    using Q_def regular_t0_eq_Hausdorff_space t0_space_Kolmogorov_quotient by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4305
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4306
  have "compactin (subtopology X Q) (Kolmogorov_quotient X ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4307
    by (metis Q_def \<open>compactin X S\<close> image_compactin quotient_imp_continuous_map quotient_map_Kolmogorov_quotient)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4308
  ultimately obtain h where conth: "continuous_map (subtopology X Q) euclidean h" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4309
              and him: "h ` topspace (subtopology X Q) \<subseteq> T" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4310
              and hg: "\<And>x. x \<in> Kolmogorov_quotient X ` S \<Longrightarrow> h x = g x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4311
    using Tietze_extension_comp_reg_aux [of "subtopology X Q" "Kolmogorov_quotient X ` S" T g]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4312
    apply (simp add: subtopology_subtopology creg contg assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4313
    using fim gf by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4314
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4315
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4316
    show "continuous_map X euclideanreal (h \<circ> Kolmogorov_quotient X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4317
      by (metis Q_def conth continuous_map_compose quotient_imp_continuous_map quotient_map_Kolmogorov_quotient)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4318
    show "(h \<circ> Kolmogorov_quotient X) ` topspace X \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4319
      using Q_def continuous_map_Kolmogorov_quotient continuous_map_image_subset_topspace him by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4320
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4321
    assume "x \<in> S" then show "(h \<circ> Kolmogorov_quotient X) x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4322
      by (simp add: gf hg)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4323
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4324
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4325
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4326
subsection\<open>Size bounds on connected or path-connected spaces\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4327
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4328
lemma connected_space_imp_card_ge_alt:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4329
  assumes "connected_space X" "completely_regular_space X" "closedin X S" "S \<noteq> {}" "S \<noteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4330
  shows "(UNIV::real set) \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4331
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4332
  have "S \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4333
    using \<open>closedin X S\<close> closedin_subset by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4334
  then obtain a where "a \<in> topspace X" "a \<notin> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4335
    using \<open>S \<noteq> topspace X\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4336
  have "(UNIV::real set) \<lesssim> {0..1::real}"
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  4337
    using eqpoll_real_subset 
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4338
    by (meson atLeastAtMost_iff eqpoll_imp_lepoll eqpoll_sym less_eq_real_def zero_less_one)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4339
  also have "\<dots> \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4340
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4341
    obtain f where contf: "continuous_map X euclidean f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4342
      and fim: "f \<in> (topspace X) \<rightarrow> {0..1::real}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4343
      and f0: "f a = 0" and f1: "f ` S \<subseteq> {1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4344
      using \<open>completely_regular_space X\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4345
      unfolding completely_regular_space_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4346
      by (metis Diff_iff \<open>a \<in> topspace X\<close> \<open>a \<notin> S\<close> \<open>closedin X S\<close> continuous_map_in_subtopology image_subset_iff_funcset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4347
    have "\<exists>y\<in>topspace X. x = f y" if "0 \<le> x" and "x \<le> 1" for x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4348
    proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4349
      have "connectedin euclidean (f ` topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4350
        using \<open>connected_space X\<close> connectedin_continuous_map_image connectedin_topspace contf by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4351
      moreover have "\<exists>y. 0 = f y \<and> y \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4352
        using \<open>a \<in> topspace X\<close> f0 by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4353
      moreover have "\<exists>y. 1 = f y \<and> y \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4354
        using \<open>S \<subseteq> topspace X\<close> \<open>S \<noteq> {}\<close> f1 by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4355
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4356
        using that by (fastforce simp: is_interval_1 simp flip: is_interval_connected_1)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4357
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4358
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4359
      unfolding lepoll_iff using atLeastAtMost_iff by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4360
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4361
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4362
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4363
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4364
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4365
lemma connected_space_imp_card_ge_gen:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4366
  assumes "connected_space X" "normal_space X" "closedin X S" "closedin X T" "S \<noteq> {}" "T \<noteq> {}" "disjnt S T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4367
  shows "(UNIV::real set) \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4368
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4369
  have "(UNIV::real set) \<lesssim> {0..1::real}"
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  4370
    by (metis atLeastAtMost_iff eqpoll_real_subset eqpoll_imp_lepoll eqpoll_sym less_le_not_le zero_less_one)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4371
  also have "\<dots>\<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4372
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4373
    obtain f where contf: "continuous_map X euclidean f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4374
       and fim: "f \<in> (topspace X) \<rightarrow> {0..1::real}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4375
       and f0: "f ` S \<subseteq> {0}" and f1: "f ` T \<subseteq> {1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4376
      using assms by (metis continuous_map_in_subtopology normal_space_iff_Urysohn image_subset_iff_funcset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4377
    have "\<exists>y\<in>topspace X. x = f y" if "0 \<le> x" and "x \<le> 1" for x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4378
    proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4379
      have "connectedin euclidean (f ` topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4380
        using \<open>connected_space X\<close> connectedin_continuous_map_image connectedin_topspace contf by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4381
      moreover have "\<exists>y. 0 = f y \<and> y \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4382
        using \<open>closedin X S\<close> \<open>S \<noteq> {}\<close> closedin_subset f0 by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4383
      moreover have "\<exists>y. 1 = f y \<and> y \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4384
        using \<open>closedin X T\<close> \<open>T \<noteq> {}\<close> closedin_subset f1 by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4385
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4386
        using that by (fastforce simp: is_interval_1 simp flip: is_interval_connected_1)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4387
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4388
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4389
      unfolding lepoll_iff using atLeastAtMost_iff by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4390
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4391
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4392
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4393
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4394
lemma connected_space_imp_card_ge:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4395
  assumes "connected_space X" "normal_space X" "t1_space X" and nosing: "\<not> (\<exists>a. topspace X \<subseteq> {a})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4396
  shows "(UNIV::real set) \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4397
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4398
  obtain a b where "a \<in> topspace X" "b \<in> topspace X" "a \<noteq> b"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4399
    by (metis nosing singletonI subset_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4400
  then have "{a} \<noteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4401
    by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4402
  with connected_space_imp_card_ge_alt assms show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4403
    by (metis \<open>a \<in> topspace X\<close> closedin_t1_singleton insert_not_empty normal_imp_completely_regular_space_A)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4404
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4405
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4406
lemma connected_space_imp_infinite_gen:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4407
   "\<lbrakk>connected_space X; t1_space X; \<nexists>a. topspace X \<subseteq> {a}\<rbrakk> \<Longrightarrow> infinite(topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4408
  by (metis connected_space_discrete_topology finite_t1_space_imp_discrete_topology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4409
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4410
lemma connected_space_imp_infinite:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4411
   "\<lbrakk>connected_space X; Hausdorff_space X; \<nexists>a. topspace X \<subseteq> {a}\<rbrakk> \<Longrightarrow> infinite(topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4412
  by (simp add: Hausdorff_imp_t1_space connected_space_imp_infinite_gen)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4413
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4414
lemma connected_space_imp_infinite_alt:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4415
  assumes "connected_space X" "regular_space X" "closedin X S" "S \<noteq> {}" "S \<noteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4416
  shows "infinite(topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4417
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4418
  have "S \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4419
    using \<open>closedin X S\<close> closedin_subset by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4420
  then obtain a where a: "a \<in> topspace X" "a \<notin> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4421
    using \<open>S \<noteq> topspace X\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4422
  have "\<exists>\<Phi>. \<forall>n. (disjnt (\<Phi> n) S \<and> a \<in> \<Phi> n \<and> openin X (\<Phi> n)) \<and> \<Phi>(Suc n) \<subset> \<Phi> n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4423
  proof (rule dependent_nat_choice)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4424
    show "\<exists>T. disjnt T S \<and> a \<in> T \<and> openin X T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4425
      by (metis Diff_iff a \<open>closedin X S\<close> closedin_def disjnt_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4426
    fix V n
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4427
    assume \<section>: "disjnt V S \<and> a \<in> V \<and> openin X V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4428
    then obtain U C where U: "openin X U" "closedin X C" "a \<in> U" "U \<subseteq> C" "C \<subseteq> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4429
      using \<open>regular_space X\<close> by (metis neighbourhood_base_of neighbourhood_base_of_closedin)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4430
    with assms have "U \<subset> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4431
      by (metis "\<section>" \<open>S \<subseteq> topspace X\<close> connected_space_clopen_in disjnt_def empty_iff inf.absorb_iff2 inf.orderE psubsetI subset_trans)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4432
    with U show "\<exists>U. (disjnt U S \<and> a \<in> U \<and> openin X U) \<and> U \<subset> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4433
      using "\<section>" disjnt_subset1 by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4434
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4435
  then obtain \<Phi> where \<Phi>: "\<And>n. disjnt (\<Phi> n) S \<and> a \<in> \<Phi> n \<and> openin X (\<Phi> n)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4436
    and \<Phi>sub: "\<And>n. \<Phi> (Suc n) \<subset> \<Phi> n" by metis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4437
  then have "decseq \<Phi>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4438
    by (simp add: decseq_SucI psubset_eq)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4439
  have "\<forall>n. \<exists>x. x \<in> \<Phi> n \<and> x \<notin> \<Phi>(Suc n)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4440
    by (meson \<Phi>sub psubsetE subsetI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4441
  then obtain f where fin: "\<And>n. f n \<in> \<Phi> n" and fout: "\<And>n. f n \<notin> \<Phi>(Suc n)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4442
    by metis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4443
  have "range f \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4444
    by (meson \<Phi> fin image_subset_iff openin_subset subset_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4445
  moreover have "inj f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4446
    by (metis Suc_le_eq \<open>decseq \<Phi>\<close> decseq_def fin fout linorder_injI subsetD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4447
  ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4448
    using infinite_iff_countable_subset by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4449
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4450
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4451
lemma path_connected_space_imp_card_ge:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4452
  assumes "path_connected_space X" "Hausdorff_space X" and nosing: "\<not> (\<exists>x. topspace X \<subseteq> {x})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4453
  shows "(UNIV::real set) \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4454
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4455
  obtain a b where "a \<in> topspace X" "b \<in> topspace X" "a \<noteq> b"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4456
    by (metis nosing singletonI subset_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4457
  then obtain \<gamma> where \<gamma>: "pathin X \<gamma>" "\<gamma> 0 = a" "\<gamma> 1 = b"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4458
    by (meson \<open>a \<in> topspace X\<close> \<open>b \<in> topspace X\<close> \<open>path_connected_space X\<close> path_connected_space_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4459
  let ?Y = "subtopology X (\<gamma> ` (topspace (subtopology euclidean {0..1})))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4460
  have "(UNIV::real set) \<lesssim>  topspace ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4461
  proof (intro compact_Hausdorff_or_regular_imp_normal_space connected_space_imp_card_ge)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4462
    show "connected_space ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4463
      using \<open>pathin X \<gamma>\<close> connectedin_def connectedin_path_image by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4464
    show "Hausdorff_space ?Y \<or> regular_space ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4465
      using Hausdorff_space_subtopology \<open>Hausdorff_space X\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4466
    show "t1_space ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4467
      using Hausdorff_imp_t1_space \<open>Hausdorff_space X\<close> t1_space_subtopology by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4468
    show "compact_space ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4469
      by (simp add: \<open>pathin X \<gamma>\<close> compact_space_subtopology compactin_path_image)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4470
    have "a \<in> topspace ?Y" "b \<in> topspace ?Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4471
      using \<gamma> pathin_subtopology by fastforce+
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4472
    with \<open>a \<noteq> b\<close> show "\<nexists>x. topspace ?Y \<subseteq> {x}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4473
      by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4474
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4475
  also have "\<dots> \<lesssim> \<gamma> ` {0..1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4476
    by (simp add: subset_imp_lepoll)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4477
  also have "\<dots> \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4478
    by (meson \<gamma> path_image_subset_topspace subset_imp_lepoll image_subset_iff_funcset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4479
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4480
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4481
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4482
lemma connected_space_imp_uncountable:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4483
  assumes "connected_space X" "regular_space X" "Hausdorff_space X" "\<not> (\<exists>a. topspace X \<subseteq> {a})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4484
  shows "\<not> countable(topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4485
proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4486
  assume coX: "countable (topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4487
  with \<open>regular_space X\<close> have "normal_space X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4488
    using countable_imp_Lindelof_space regular_Lindelof_imp_normal_space by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4489
  then have "(UNIV::real set) \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4490
    by (simp add: Hausdorff_imp_t1_space assms connected_space_imp_card_ge)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4491
  with coX show False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4492
    using countable_lepoll uncountable_UNIV_real by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4493
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4494
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4495
lemma path_connected_space_imp_uncountable:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4496
  assumes "path_connected_space X" "t1_space X" and nosing: "\<not> (\<exists>a. topspace X \<subseteq> {a})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4497
  shows "\<not> countable(topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4498
proof 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4499
  assume coX: "countable (topspace X)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4500
  obtain a b where "a \<in> topspace X" "b \<in> topspace X" "a \<noteq> b"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4501
    by (metis nosing singletonI subset_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4502
  then obtain \<gamma> where "pathin X \<gamma>" "\<gamma> 0 = a" "\<gamma> 1 = b"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4503
    by (meson \<open>a \<in> topspace X\<close> \<open>b \<in> topspace X\<close> \<open>path_connected_space X\<close> path_connected_space_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4504
  then have "\<gamma> ` {0..1} \<lesssim> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4505
    by (meson path_image_subset_topspace subset_imp_lepoll image_subset_iff_funcset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4506
  define \<A> where "\<A> \<equiv> ((\<lambda>a. {x \<in> {0..1}. \<gamma> x \<in> {a}}) ` topspace X) - {{}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4507
  have \<A>01: "\<A> = {{0..1}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4508
  proof (rule real_Sierpinski_lemma)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4509
    show "countable \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4510
      using \<A>_def coX by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4511
    show "disjoint \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4512
      by (auto simp: \<A>_def disjnt_iff pairwise_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4513
    show "\<Union>\<A> = {0..1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4514
      using \<open>pathin X \<gamma>\<close> path_image_subset_topspace by (fastforce simp: \<A>_def Bex_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4515
    fix C
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4516
    assume "C \<in> \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4517
    then obtain a where "a \<in> topspace X" and C: "C = {x \<in> {0..1}. \<gamma> x \<in> {a}}" "C \<noteq> {}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4518
      by (auto simp: \<A>_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4519
    then have "closedin X {a}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4520
      by (meson \<open>t1_space X\<close> closedin_t1_singleton)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4521
    then have "closedin (top_of_set {0..1}) C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4522
      using C \<open>pathin X \<gamma>\<close> closedin_continuous_map_preimage pathin_def by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4523
    then show "closed C \<and> C \<noteq> {}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4524
      using C closedin_closed_trans by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4525
  qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4526
  then have "{0..1} \<in> \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4527
    by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4528
  then have "\<exists>a \<in> topspace X. {0..1} \<subseteq> {x. \<gamma> x = a}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4529
    using \<A>_def image_iff by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4530
  then show False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4531
    using \<open>\<gamma> 0 = a\<close> \<open>\<gamma> 1 = b\<close> \<open>a \<noteq> b\<close> atLeastAtMost_iff zero_less_one_class.zero_le_one by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4532
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4533
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4534
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4535
subsection\<open>Lavrentiev extension etc\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4536
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4537
lemma (in Metric_space) convergent_eq_zero_oscillation_gen:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4538
  assumes "mcomplete" and fim: "f \<in> (topspace X \<inter> S) \<rightarrow> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4539
  shows "(\<exists>l. limitin mtopology f l (atin_within X a S)) \<longleftrightarrow>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4540
         M \<noteq> {} \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4541
         (a \<in> topspace X
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4542
            \<longrightarrow> (\<forall>\<epsilon>>0. \<exists>U. openin X U \<and> a \<in> U \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4543
                           (\<forall>x \<in> (S \<inter> U) - {a}. \<forall>y \<in> (S \<inter> U) - {a}. d (f x) (f y) < \<epsilon>)))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4544
proof (cases "M = {}")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4545
  case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4546
  with limitin_mspace show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4547
    by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4548
next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4549
  case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4550
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4551
  proof (cases "a \<in> topspace X")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4552
    case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4553
    let ?R = "\<forall>\<epsilon>>0. \<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4554
    show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4555
    proof (cases "a \<in> X derived_set_of S")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4556
      case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4557
      have ?R
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4558
        if "limitin mtopology f l (atin_within X a S)" for l
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4559
      proof (intro strip)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4560
        fix \<epsilon>::real
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4561
        assume "\<epsilon>>0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4562
        with that \<open>a \<in> topspace X\<close> 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4563
        obtain U where U: "openin X U" "a \<in> U" "l \<in> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4564
          and Uless: "\<forall>x\<in>U - {a}. x \<in> S \<longrightarrow> f x \<in> M \<and> d (f x) l < \<epsilon>/2"
78750
f229090cb8a3 atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  4565
          unfolding limitin_metric eventually_atin_within by (metis Diff_iff insertI1  half_gt_zero [OF \<open>\<epsilon>>0\<close>])
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4566
        show "\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4567
        proof (intro exI strip conjI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4568
          fix x y
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4569
          assume x: "x \<in> S \<inter> U - {a}" and y: "y \<in> S \<inter> U - {a}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4570
          then have "d (f x) l < \<epsilon>/2" "d (f y) l < \<epsilon>/2" "f x \<in> M" "f y \<in> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4571
            using Uless by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4572
          then show "d (f x) (f y) < \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4573
            using triangle' \<open>l \<in> M\<close> by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4574
        qed (auto simp add: U)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4575
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4576
      moreover have "\<exists>l. limitin mtopology f l (atin_within X a S)" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4577
        if R [rule_format]: ?R
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4578
      proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4579
        define F where "F \<equiv> \<lambda>U. mtopology closure_of f ` (S \<inter> U - {a})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4580
        define \<C> where "\<C> \<equiv> F ` {U. openin X U \<and> a \<in> U}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4581
        have \<C>_clo: "\<forall>C \<in> \<C>. closedin mtopology C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4582
          by (force simp add: \<C>_def F_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4583
        moreover have sub_mcball: "\<exists>C a. C \<in> \<C> \<and> C \<subseteq> mcball a \<epsilon>" if "\<epsilon>>0" for \<epsilon>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4584
        proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4585
          obtain U where U: "openin X U" "a \<in> U" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4586
            and Uless: "\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4587
            using R [OF \<open>\<epsilon>>0\<close>] by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4588
          then obtain b where b: "b \<noteq> a" "b \<in> S" "b \<in> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4589
            using True by (auto simp add: in_derived_set_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4590
          have "U \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4591
            by (simp add: U(1) openin_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4592
          have "f b \<in> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4593
            using b \<open>openin X U\<close> by (metis image_subset_iff_funcset Int_iff fim image_eqI openin_subset subsetD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4594
          moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4595
          have "mtopology closure_of f ` ((S \<inter> U) - {a}) \<subseteq> mcball (f b) \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4596
          proof (rule closure_of_minimal)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4597
            have "f y \<in> M" if "y \<in> S" and "y \<in> U" for y
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4598
              using \<open>U \<subseteq> topspace X\<close> fim that by (auto simp: Pi_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4599
            moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4600
            have "d (f b) (f y) \<le> \<epsilon>" if "y \<in> S" "y \<in> U" "y \<noteq> a" for y
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4601
              using that Uless b by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4602
            ultimately show "f ` (S \<inter> U - {a}) \<subseteq> mcball (f b) \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4603
              by (force simp: \<open>f b \<in> M\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4604
          qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4605
          ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4606
            using U by (auto simp add: \<C>_def F_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4607
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4608
        moreover have "\<Inter>\<F> \<noteq> {}" if "finite \<F>" "\<F> \<subseteq> \<C>" for \<F>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4609
        proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4610
          obtain \<G> where sub: "\<G> \<subseteq> {U. openin X U \<and> a \<in> U}" and eq: "\<F> = F ` \<G>" and "finite \<G>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4611
            by (metis (no_types, lifting) \<C>_def \<open>\<F> \<subseteq> \<C>\<close> \<open>finite \<F>\<close> finite_subset_image)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4612
          then have "U \<subseteq> topspace X" if "U \<in> \<G>" for U
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4613
            using openin_subset that by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4614
          then have "T \<subseteq> mtopology closure_of T" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4615
            if "T \<in> (\<lambda>U. f ` (S \<inter> U - {a})) ` \<G>" for T
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4616
            using that fim by (fastforce simp add: intro!: closure_of_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4617
          moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4618
          have ain: "a \<in> \<Inter> (insert (topspace X) \<G>)" "openin X (\<Inter> (insert (topspace X) \<G>))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4619
            using True in_derived_set_of sub \<open>finite \<G>\<close> by (fastforce intro!: openin_Inter)+
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4620
          then obtain y where "y \<noteq> a" "y \<in> S" and y: "y \<in> \<Inter> (insert (topspace X) \<G>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4621
            by (meson \<open>a \<in> X derived_set_of S\<close> sub in_derived_set_of) 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4622
          then have "f y \<in> \<Inter>\<F>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4623
            using eq that ain fim by (auto simp add: F_def image_subset_iff in_closure_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4624
          then show ?thesis by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4625
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4626
        ultimately have "\<Inter>\<C> \<noteq> {}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4627
          using \<open>mcomplete\<close> mcomplete_fip by metis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4628
        then obtain b where "b \<in> \<Inter>\<C>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4629
          by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4630
        then have "b \<in> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4631
          using sub_mcball \<C>_clo mbounded_alt_pos mbounded_empty metric_closedin_iff_sequentially_closed by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4632
        have "limitin mtopology f b (atin_within X a S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4633
        proof (clarsimp simp: limitin_metric \<open>b \<in> M\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4634
          fix \<epsilon> :: real
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4635
          assume "\<epsilon> > 0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4636
          then obtain U where U: "openin X U" "a \<in> U" and subU: "U \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4637
            and Uless: "\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>/2"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4638
            by (metis R half_gt_zero openin_subset) 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4639
          then obtain x where x: "x \<in> S" "x \<in> U" "x \<noteq> a" and fx: "f x \<in> mball b (\<epsilon>/2)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4640
            using \<open>b \<in> \<Inter>\<C>\<close> 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4641
            apply (simp add: \<C>_def F_def closure_of_def del: divide_const_simps)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4642
            by (metis Diff_iff Int_iff centre_in_mball_iff in_mball openin_mball singletonI zero_less_numeral)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4643
          moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4644
          have "d (f y) b < \<epsilon>" if "y \<in> U" "y \<noteq> a" "y \<in> S" for y
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4645
          proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4646
            have "d (f x) (f y) < \<epsilon>/2"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4647
              using Uless that x by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4648
            moreover have "d b (f x)  < \<epsilon>/2"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4649
              using fx by simp
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4650
            ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4651
              using triangle [of b "f x" "f y"] subU that \<open>b \<in> M\<close> commute fim fx by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4652
          qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4653
          ultimately show "\<forall>\<^sub>F x in atin_within X a S. f x \<in> M \<and> d (f x) b < \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4654
            using fim U
78750
f229090cb8a3 atin now an abbreviation for atin_within, which has been moved to Abstract_Limits
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  4655
            apply (simp add: eventually_atin_within del: divide_const_simps flip: image_subset_iff_funcset)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4656
            by (smt (verit, del_insts) Diff_iff Int_iff imageI insertI1 openin_subset subsetD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4657
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4658
        then show ?thesis ..
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4659
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4660
      ultimately
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4661
      show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4662
        by (meson True \<open>M \<noteq> {}\<close> in_derived_set_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4663
    next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4664
      case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4665
      have "(\<exists>l. limitin mtopology f l (atin_within X a S))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4666
        by (metis \<open>M \<noteq> {}\<close> False derived_set_of_trivial_limit equals0I limitin_trivial topspace_mtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4667
      moreover have "\<forall>e>0. \<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < e)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4668
        by (metis Diff_iff False IntE True in_derived_set_of insert_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4669
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4670
        using limitin_mspace by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4671
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4672
  next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4673
    case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4674
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4675
      by (metis derived_set_of_trivial_limit ex_in_conv in_derived_set_of limitin_mspace limitin_trivial topspace_mtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4676
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4677
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4678
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4679
text \<open>The HOL Light proof uses some ugly tricks to share common parts of what are two separate proofs for the two cases\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4680
lemma (in Metric_space) gdelta_in_points_of_convergence_within:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4681
  assumes "mcomplete"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4682
    and f: "continuous_map (subtopology X S) mtopology f \<or> t1_space X \<and> f \<in> S \<rightarrow> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4683
  shows "gdelta_in X {x \<in> topspace X. \<exists>l. limitin mtopology f l (atin_within X x S)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4684
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4685
  have fim: "f \<in> (topspace X \<inter> S) \<rightarrow> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4686
    using continuous_map_image_subset_topspace f by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4687
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4688
  proof (cases "M={}")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4689
    case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4690
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4691
      by (smt (verit) Collect_cong empty_def empty_iff gdelta_in_empty limitin_mspace)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4692
  next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4693
    case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4694
    define A where "A \<equiv> {a \<in> topspace X. \<forall>\<epsilon>>0. \<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4695
    have "gdelta_in X A"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4696
      using f 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4697
    proof (elim disjE conjE)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4698
      assume cm: "continuous_map (subtopology X S) mtopology f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4699
      define C where "C \<equiv> \<lambda>r. \<Union>{U. openin X U \<and> (\<forall>x \<in> S\<inter>U. \<forall>y \<in> S\<inter>U. d (f x) (f y) < r)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4700
      define B where "B \<equiv> (\<Inter>n. C(inverse(Suc n)))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4701
      define D where "D \<equiv> (\<Inter> (C ` {0<..}))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4702
      have "D=B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4703
        unfolding B_def C_def D_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4704
        apply (intro Inter_eq_Inter_inverse_Suc Sup_subset_mono)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4705
        by (smt (verit, ccfv_threshold) Collect_mono_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4706
      have "B \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4707
        using openin_subset by (force simp add: B_def C_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4708
      have "(countable intersection_of openin X) B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4709
        unfolding B_def C_def 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4710
        by (intro relative_to_inc countable_intersection_of_Inter countable_intersection_of_inc) auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4711
      then have "gdelta_in X B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4712
        unfolding gdelta_in_def by (intro relative_to_subset_inc \<open>B \<subseteq> topspace X\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4713
      moreover have "A=D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4714
      proof (intro equalityI subsetI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4715
        fix a
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4716
        assume x: "a \<in> A"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4717
        then have "a \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4718
          using A_def by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4719
        show "a \<in> D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4720
        proof (clarsimp simp: D_def C_def \<open>a \<in> topspace X\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4721
          fix \<epsilon>::real assume "\<epsilon> > 0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4722
          then obtain U where "openin X U" "a \<in> U" and U: "(\<forall>x\<in>S \<inter> U - {a}. \<forall>y\<in>S \<inter> U - {a}. d (f x) (f y) < \<epsilon>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4723
            using x by (force simp: A_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4724
          show "\<exists>T. openin X T \<and> (\<forall>x\<in>S \<inter> T. \<forall>y\<in>S \<inter> T. d (f x) (f y) < \<epsilon>) \<and> a \<in> T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4725
          proof (cases "a \<in> S")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4726
            case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4727
            then obtain V where "openin X V" "a \<in> V" and V: "\<forall>x. x \<in> S \<and> x \<in> V \<longrightarrow> f a \<in> M \<and> f x \<in> M \<and> d (f a) (f x) < \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4728
              using \<open>a \<in> topspace X\<close> \<open>\<epsilon> > 0\<close> cm
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4729
              by (force simp add: continuous_map_to_metric openin_subtopology_alt Ball_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4730
            show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4731
            proof (intro exI conjI strip)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4732
              show "openin X (U \<inter> V)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4733
                using \<open>openin X U\<close> \<open>openin X V\<close> by blast 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4734
              show "a \<in> U \<inter> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4735
                using \<open>a \<in> U\<close> \<open>a \<in> V\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4736
              show "\<And>x y. \<lbrakk>x \<in> S \<inter> (U \<inter> V); y \<in> S \<inter> (U \<inter> V)\<rbrakk> \<Longrightarrow> d (f x) (f y) < \<epsilon>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4737
                by (metis DiffI Int_iff U V commute singletonD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4738
            qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4739
          next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4740
            case False then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4741
              using U \<open>a \<in> U\<close> \<open>openin X U\<close> by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4742
          qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4743
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4744
      next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4745
        fix x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4746
        assume x: "x \<in> D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4747
        then have "x \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4748
          using \<open>B \<subseteq> topspace X\<close> \<open>D=B\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4749
        with x show "x \<in> A"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4750
          apply (clarsimp simp: D_def C_def A_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4751
          by (meson DiffD1 greaterThan_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4752
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4753
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4754
        by (simp add: \<open>D=B\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4755
    next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4756
      assume "t1_space X" "f \<in> S \<rightarrow> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4757
      define C where "C \<equiv> \<lambda>r. \<Union>{U. openin X U \<and> 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4758
                           (\<exists>b \<in> topspace X. \<forall>x \<in> S\<inter>U - {b}. \<forall>y \<in> S\<inter>U - {b}. d (f x) (f y) < r)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4759
      define B where "B \<equiv> (\<Inter>n. C(inverse(Suc n)))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4760
      define D where "D \<equiv> (\<Inter> (C ` {0<..}))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4761
      have "D=B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4762
        unfolding B_def C_def D_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4763
        apply (intro Inter_eq_Inter_inverse_Suc Sup_subset_mono)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4764
        by (smt (verit, ccfv_threshold) Collect_mono_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4765
      have "B \<subseteq> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4766
        using openin_subset by (force simp add: B_def C_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4767
      have "(countable intersection_of openin X) B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4768
        unfolding B_def C_def 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4769
        by (intro relative_to_inc countable_intersection_of_Inter countable_intersection_of_inc) auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4770
      then have "gdelta_in X B"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4771
        unfolding gdelta_in_def by (intro relative_to_subset_inc \<open>B \<subseteq> topspace X\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4772
      moreover have "A=D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4773
      proof (intro equalityI subsetI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4774
        fix x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4775
        assume x: "x \<in> D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4776
        then have "x \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4777
          using \<open>B \<subseteq> topspace X\<close> \<open>D=B\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4778
        show "x \<in> A"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4779
        proof (clarsimp simp: A_def \<open>x \<in> topspace X\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4780
          fix \<epsilon> :: real
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4781
          assume "\<epsilon>>0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4782
          then obtain U b where "openin X U" "b \<in> topspace X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4783
                  and U: "\<forall>x\<in>S \<inter> U - {b}. \<forall>y\<in>S \<inter> U - {b}. d (f x) (f y) < \<epsilon>" and "x \<in> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4784
            using x by (auto simp: D_def C_def A_def Ball_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4785
          then have "openin X (U-{b})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4786
            by (meson \<open>t1_space X\<close> t1_space_openin_delete_alt)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4787
          then show "\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>xa\<in>S \<inter> U - {x}. \<forall>y\<in>S \<inter> U - {x}. d (f xa) (f y) < \<epsilon>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4788
            using U \<open>openin X U\<close> \<open>x \<in> U\<close> by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4789
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4790
      next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4791
        show "\<And>x. x \<in> A \<Longrightarrow> x \<in> D"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4792
          unfolding A_def D_def C_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4793
          by clarsimp meson
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4794
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4795
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4796
        by (simp add: \<open>D=B\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4797
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4798
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4799
      by (simp add: A_def convergent_eq_zero_oscillation_gen False fim \<open>mcomplete\<close> cong: conj_cong)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4800
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4801
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4802
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4803
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4804
lemma gdelta_in_points_of_convergence_within:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4805
  assumes Y: "completely_metrizable_space Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4806
    and f: "continuous_map (subtopology X S) Y f \<or> t1_space X \<and> f \<in> S \<rightarrow> topspace Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4807
  shows "gdelta_in X {x \<in> topspace X. \<exists>l. limitin Y f l (atin_within X x S)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4808
  using assms
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4809
  unfolding completely_metrizable_space_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4810
  using Metric_space.gdelta_in_points_of_convergence_within Metric_space.topspace_mtopology by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4811
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4812
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4813
lemma Lavrentiev_extension_gen:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4814
  assumes "S \<subseteq> topspace X" and Y: "completely_metrizable_space Y" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4815
    and contf: "continuous_map (subtopology X S) Y f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4816
  obtains U g where "gdelta_in X U" "S \<subseteq> U" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4817
            "continuous_map (subtopology X (X closure_of S \<inter> U)) Y g" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4818
             "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4819
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4820
  define U where "U \<equiv> {x \<in> topspace X. \<exists>l. limitin Y f l (atin_within X x S)}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4821
  have "S \<subseteq> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4822
    using that contf limit_continuous_map_within subsetD [OF \<open>S \<subseteq> topspace X\<close>] 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4823
    by (fastforce simp: U_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4824
  then have "S \<subseteq> X closure_of S \<inter> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4825
    by (simp add: \<open>S \<subseteq> topspace X\<close> closure_of_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4826
  moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4827
  have "\<And>t. t \<in> X closure_of S \<inter> U - S \<Longrightarrow> \<exists>l. limitin Y f l (atin_within X t S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4828
    using U_def by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4829
  moreover have "regular_space Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4830
    by (simp add: Y completely_metrizable_imp_metrizable_space metrizable_imp_regular_space)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4831
  ultimately
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4832
  obtain g where g: "continuous_map (subtopology X (X closure_of S \<inter> U)) Y g" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4833
    and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4834
    using continuous_map_extension_pointwise_alt assms by blast 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4835
  show thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4836
  proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4837
    show "gdelta_in X U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4838
      by (simp add: U_def Y contf gdelta_in_points_of_convergence_within)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4839
    show "continuous_map (subtopology X (X closure_of S \<inter> U)) Y g"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4840
      by (simp add: g)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4841
  qed (use \<open>S \<subseteq> U\<close> gf in auto)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4842
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4843
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4844
lemma Lavrentiev_extension:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4845
  assumes "S \<subseteq> topspace X" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4846
    and X: "metrizable_space X \<or> topspace X \<subseteq> X closure_of S" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4847
    and Y: "completely_metrizable_space Y" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4848
    and contf: "continuous_map (subtopology X S) Y f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4849
  obtains U g where "gdelta_in X U" "S \<subseteq> U" "U \<subseteq> X closure_of S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4850
            "continuous_map (subtopology X U) Y g"  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4851
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4852
  obtain U g where "gdelta_in X U" "S \<subseteq> U" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4853
    and contg: "continuous_map (subtopology X (X closure_of S \<inter> U)) Y g" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4854
    and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4855
    using Lavrentiev_extension_gen Y assms(1) contf by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4856
  define V where "V \<equiv> X closure_of S \<inter> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4857
  show thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4858
  proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4859
    show "gdelta_in X V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4860
      by (metis V_def X \<open>gdelta_in X U\<close> closed_imp_gdelta_in closedin_closure_of closure_of_subset_topspace gdelta_in_Int gdelta_in_topspace subset_antisym)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4861
    show "S \<subseteq> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4862
      by (simp add: V_def \<open>S \<subseteq> U\<close> assms(1) closure_of_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4863
    show "continuous_map (subtopology X V) Y g"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4864
      by (simp add: V_def contg)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4865
  qed (auto simp: gf V_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4866
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4867
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4868
subsection\<open>Embedding in products and hence more about completely metrizable spaces\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4869
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4870
lemma (in Metric_space) gdelta_homeomorphic_space_closedin_product:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4871
  assumes S: "\<And>i. i \<in> I \<Longrightarrow> openin mtopology (S i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4872
  obtains T where "closedin (prod_topology mtopology (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4873
                  "subtopology mtopology (\<Inter>i \<in> I. S i) homeomorphic_space
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4874
                   subtopology (prod_topology mtopology (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4875
proof (cases "I={}")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4876
  case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4877
  then have top: "topspace (prod_topology mtopology (powertop_real I)) = (M \<times> {(\<lambda>x. undefined)})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4878
    by simp
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4879
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4880
  proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4881
    show "closedin (prod_topology mtopology (powertop_real I)) (M \<times> {(\<lambda>x. undefined)})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4882
      by (metis top closedin_topspace)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4883
    have "subtopology mtopology (\<Inter>(S ` I)) homeomorphic_space mtopology"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4884
      by (simp add: True product_topology_empty_discrete)
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  4885
    also have "\<dots> homeomorphic_space (prod_topology mtopology (discrete_topology {\<lambda>x. undefined}))"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  4886
      by (meson homeomorphic_space_sym prod_topology_homeomorphic_space_left)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4887
    finally
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4888
    show "subtopology mtopology (\<Inter>(S ` I)) homeomorphic_space subtopology (prod_topology mtopology (powertop_real I)) (M \<times> {(\<lambda>x. undefined)})"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  4889
      by (smt (verit, ccfv_SIG) True product_topology_empty_discrete subtopology_topspace top)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4890
  qed   
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4891
next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4892
  case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4893
  have SM: "\<And>i. i \<in> I \<Longrightarrow> S i \<subseteq> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4894
    using assms openin_mtopology by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4895
  then have "(\<Inter>i \<in> I. S i) \<subseteq> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4896
    using False by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4897
  define dd where "dd \<equiv> \<lambda>i. if i\<notin>I \<or> S i = M then \<lambda>u. 1 else (\<lambda>u. INF x \<in> M - S i. d u x)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4898
  have [simp]: "bdd_below (d u ` A)" for u A
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4899
    by (meson bdd_belowI2 nonneg)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4900
  have cont_dd: "continuous_map (subtopology mtopology (S i)) euclidean (dd i)" if "i \<in> I" for i
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4901
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4902
    have "dist (Inf (d x ` (M - S i))) (Inf (d y ` (M - S i))) \<le> d x y" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4903
      if "x \<in> S i" "x \<in> M" "y \<in> S i" "y \<in> M" "S i \<noteq> M" for x y
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4904
    proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4905
      have [simp]: "\<not> M \<subseteq> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4906
        using SM \<open>S i \<noteq> M\<close> \<open>i \<in> I\<close> by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4907
      have "\<And>u. \<lbrakk>u \<in> M; u \<notin> S i\<rbrakk> \<Longrightarrow> Inf (d x ` (M - S i)) \<le> d x y + d y u"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4908
        apply (clarsimp simp add: cInf_le_iff_less)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4909
        by (smt (verit) DiffI triangle \<open>x \<in> M\<close> \<open>y \<in> M\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4910
      then have "Inf (d x ` (M - S i)) - d x y \<le> Inf (d y ` (M - S i))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4911
        by (force simp add: le_cInf_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4912
      moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4913
      have "\<And>u. \<lbrakk>u \<in> M; u \<notin> S i\<rbrakk> \<Longrightarrow> Inf (d y ` (M - S i)) \<le> d x u + d x y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4914
        apply (clarsimp simp add: cInf_le_iff_less)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4915
        by (smt (verit) DiffI triangle'' \<open>x \<in> M\<close> \<open>y \<in> M\<close>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4916
      then have "Inf (d y ` (M - S i)) - d x y \<le> Inf (d x ` (M - S i))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4917
        by (force simp add: le_cInf_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4918
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4919
        by (simp add: dist_real_def abs_le_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4920
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4921
    then have *: "Lipschitz_continuous_map (submetric Self (S i)) euclidean_metric (\<lambda>u. Inf (d u ` (M - S i)))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4922
      unfolding Lipschitz_continuous_map_def by (force intro!: exI [where x=1])
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4923
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4924
      using Lipschitz_continuous_imp_continuous_map [OF *]
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4925
      by (simp add: dd_def Self_def mtopology_of_submetric )
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4926
  qed 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4927
  have dd_pos: "0 < dd i x" if "x \<in> S i" for i x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4928
  proof (clarsimp simp add: dd_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4929
    assume "i \<in> I" and "S i \<noteq> M"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4930
    have opeS: "openin mtopology (S i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4931
      by (simp add: \<open>i \<in> I\<close> assms)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4932
    then obtain r where "r>0" and r: "\<And>y. \<lbrakk>y \<in> M; d x y < r\<rbrakk> \<Longrightarrow> y \<in> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4933
      by (meson \<open>x \<in> S i\<close> in_mball openin_mtopology subsetD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4934
    then have "\<And>y. y \<in> M - S i \<Longrightarrow> d x y \<ge> r"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4935
      by (meson Diff_iff linorder_not_le)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4936
    then have "Inf (d x ` (M - S i)) \<ge> r"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4937
      by (meson Diff_eq_empty_iff SM \<open>S i \<noteq> M\<close> \<open>i \<in> I\<close> cINF_greatest set_eq_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4938
    with \<open>r>0\<close> show "0 < Inf (d x ` (M - S i))" by simp
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4939
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4940
  define f where "f \<equiv> \<lambda>x. (x, \<lambda>i\<in>I. inverse(dd i x))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4941
  define T where "T \<equiv> f ` (\<Inter>i \<in> I. S i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4942
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4943
  proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4944
    show "closedin (prod_topology mtopology (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4945
      unfolding closure_of_subset_eq [symmetric]
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4946
    proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4947
      show "T \<subseteq> topspace (prod_topology mtopology (powertop_real I))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4948
        using False SM by (auto simp: T_def f_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4949
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4950
      have "(x,ds) \<in> T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4951
        if \<section>: "\<And>U. \<lbrakk>(x, ds) \<in> U; openin (prod_topology mtopology (powertop_real I)) U\<rbrakk> \<Longrightarrow> \<exists>y\<in>T. y \<in> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4952
          and "x \<in> M" and ds: "ds \<in> I \<rightarrow>\<^sub>E UNIV" for x ds
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4953
      proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4954
        have ope: "\<exists>x. x \<in> \<Inter>(S ` I) \<and> f x \<in> U \<times> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4955
          if "x \<in> U" and "ds \<in> V" and "openin mtopology U" and "openin (powertop_real I) V" for U V
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4956
          using \<section> [of "U\<times>V"] that by (force simp add: T_def openin_prod_Times_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4957
        have x_in_INT: "x \<in> \<Inter>(S ` I)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4958
        proof clarify
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4959
          fix i
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4960
          assume "i \<in> I"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4961
          show "x \<in> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4962
          proof (rule ccontr)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4963
            assume "x \<notin> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4964
            have "openin (powertop_real I) {z \<in> topspace (powertop_real I). z i \<in> {ds i - 1 <..< ds i + 1}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4965
            proof (rule openin_continuous_map_preimage)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4966
              show "continuous_map (powertop_real I) euclidean (\<lambda>x. x i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4967
                by (metis \<open>i \<in> I\<close> continuous_map_product_projection)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4968
            qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4969
            then obtain y where "y \<in> S i" "y \<in> M" and dxy: "d x y < inverse (\<bar>ds i\<bar> + 1)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4970
                          and intvl: "inverse (dd i y) \<in> {ds i - 1 <..< ds i + 1}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4971
              using ope [of "mball x (inverse(abs(ds i) + 1))" "{z \<in> topspace(powertop_real I). z i \<in> {ds i - 1<..<ds i + 1}}"]
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4972
                    \<open>x \<in> M\<close> \<open>ds \<in> I \<rightarrow>\<^sub>E UNIV\<close> \<open>i \<in> I\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4973
              by (fastforce simp add: f_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4974
            have "\<not> M \<subseteq> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4975
              using \<open>x \<notin> S i\<close> \<open>x \<in> M\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4976
            have "inverse (\<bar>ds i\<bar> + 1) \<le> dd i y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4977
              using intvl \<open>y \<in> S i\<close> dd_pos [of y i]
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4978
              by (smt (verit, ccfv_threshold) greaterThanLessThan_iff inverse_inverse_eq le_imp_inverse_le)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4979
            also have "\<dots> \<le> d x y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4980
              using \<open>i \<in> I\<close> \<open>\<not> M \<subseteq> S i\<close> \<open>x \<notin> S i\<close> \<open>x \<in> M\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4981
              apply (simp add: dd_def cInf_le_iff_less)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4982
              using commute by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4983
            finally show False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4984
              using dxy by linarith
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4985
          qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4986
        qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4987
        moreover have "ds = (\<lambda>i\<in>I. inverse (dd i x))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4988
        proof (rule PiE_ext [OF ds])
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4989
          fix i
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4990
          assume "i \<in> I"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4991
          define e where "e \<equiv> \<bar>ds i - inverse (dd i x)\<bar>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4992
          { assume con: "e > 0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4993
            have "continuous_map (subtopology mtopology (S i)) euclidean (\<lambda>x. inverse (dd i x))" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4994
              using dd_pos cont_dd \<open>i \<in> I\<close> 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4995
              by (fastforce simp:  intro!: continuous_map_real_inverse)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4996
             then have "openin (subtopology mtopology (S i))
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4997
                         {z \<in> topspace (subtopology mtopology (S i)). 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4998
                          inverse (dd i z) \<in> {inverse(dd i x) - e/2<..<inverse(dd i x) + e/2}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4999
               using openin_continuous_map_preimage open_greaterThanLessThan open_openin by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5000
             then obtain U where "openin mtopology U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5001
                  and U: "{z \<in> S i. inverse (dd i x) - e/2 < inverse (dd i z) \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5002
                           inverse (dd i z) < inverse (dd i x) + e/2}
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5003
                         = U \<inter> S i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5004
               using SM \<open>i \<in> I\<close>  by (auto simp: openin_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5005
             have "x \<in> U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5006
               using U x_in_INT \<open>i \<in> I\<close> con by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5007
             have "ds \<in> {z \<in> topspace (powertop_real I). z i \<in> {ds i - e / 2<..<ds i + e/2}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5008
               by (simp add: con ds)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5009
             moreover
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5010
             have  "openin (powertop_real I) {z \<in> topspace (powertop_real I). z i \<in> {ds i - e / 2<..<ds i + e/2}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5011
             proof (rule openin_continuous_map_preimage)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5012
               show "continuous_map (powertop_real I) euclidean (\<lambda>x. x i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5013
                 by (metis \<open>i \<in> I\<close> continuous_map_product_projection)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5014
             qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5015
             ultimately obtain y where "y \<in> \<Inter>(S ` I) \<and> f y \<in> U \<times> {z \<in> topspace (powertop_real I). z i \<in> {ds i - e / 2<..<ds i + e/2}}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5016
               using ope \<open>x \<in> U\<close> \<open>openin mtopology U\<close> \<open>x \<in> U\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5017
               by presburger 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5018
             with \<open>i \<in> I\<close> U 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5019
             have False unfolding set_eq_iff f_def e_def by simp (smt (verit) field_sum_of_halves)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5020
          }
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5021
          then show "ds i = (\<lambda>i\<in>I. inverse (dd i x)) i"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5022
            using \<open>i \<in> I\<close> by (force simp: e_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5023
        qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5024
        ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5025
          by (auto simp: T_def f_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5026
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5027
      then show "prod_topology mtopology (powertop_real I) closure_of T \<subseteq> T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5028
        by (auto simp: closure_of_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5029
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5030
    have eq: "(\<Inter>(S ` I) \<times> (I \<rightarrow>\<^sub>E UNIV) \<inter> f ` (M \<inter> \<Inter>(S ` I))) = (f ` \<Inter>(S ` I))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5031
      using False SM by (force simp: f_def image_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5032
    have "continuous_map (subtopology mtopology (\<Inter>(S ` I))) euclidean (dd i)" if "i \<in> I" for i
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5033
      by (meson INT_lower cont_dd continuous_map_from_subtopology_mono that)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5034
    then have "continuous_map (subtopology mtopology (\<Inter>(S ` I))) (powertop_real I) (\<lambda>x. \<lambda>i\<in>I. inverse (dd i x))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5035
      using dd_pos by (fastforce simp: continuous_map_componentwise intro!: continuous_map_real_inverse)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5036
    then have "embedding_map (subtopology mtopology (\<Inter>(S ` I))) (prod_topology (subtopology mtopology (\<Inter>(S ` I))) (powertop_real I)) f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5037
      by (simp add: embedding_map_graph f_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5038
    moreover have "subtopology (prod_topology (subtopology mtopology (\<Inter>(S ` I))) (powertop_real I))
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5039
                     (f ` topspace (subtopology mtopology (\<Inter>(S ` I)))) =
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5040
                   subtopology (prod_topology mtopology (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5041
      by (simp add: prod_topology_subtopology subtopology_subtopology T_def eq)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5042
    ultimately
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5043
    show "subtopology mtopology (\<Inter>(S ` I)) homeomorphic_space subtopology (prod_topology mtopology (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5044
      by (metis embedding_map_imp_homeomorphic_space)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5045
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5046
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5047
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5048
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5049
lemma gdelta_homeomorphic_space_closedin_product:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5050
  assumes "metrizable_space X" and "\<And>i. i \<in> I \<Longrightarrow> openin X (S i)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5051
  obtains T where "closedin (prod_topology X (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5052
                  "subtopology X (\<Inter>i \<in> I. S i) homeomorphic_space
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5053
                   subtopology (prod_topology X (powertop_real I)) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5054
  using Metric_space.gdelta_homeomorphic_space_closedin_product
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5055
  by (metis assms metrizable_space_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5056
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5057
lemma open_homeomorphic_space_closedin_product:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5058
  assumes "metrizable_space X" and "openin X S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5059
  obtains T where "closedin (prod_topology X euclideanreal) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5060
    "subtopology X S homeomorphic_space
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5061
                   subtopology (prod_topology X euclideanreal) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5062
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5063
  obtain T where cloT: "closedin (prod_topology X (powertop_real {()})) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5064
    and homT: "subtopology X S homeomorphic_space
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5065
                   subtopology (prod_topology X (powertop_real {()})) T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5066
    using gdelta_homeomorphic_space_closedin_product [of X "{()}" "\<lambda>i. S"] assms
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5067
    by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5068
  have "prod_topology X (powertop_real {()}) homeomorphic_space prod_topology X euclideanreal"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5069
    by (meson homeomorphic_space_prod_topology homeomorphic_space_refl homeomorphic_space_singleton_product)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5070
  then obtain f where f: "homeomorphic_map (prod_topology X (powertop_real {()})) (prod_topology X euclideanreal) f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5071
    unfolding homeomorphic_space by metis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5072
  show thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5073
  proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5074
    show "closedin (prod_topology X euclideanreal) (f ` T)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5075
      using cloT f homeomorphic_map_closedness_eq by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5076
    moreover have "T = topspace (subtopology (prod_topology X (powertop_real {()})) T)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5077
      by (metis cloT closedin_subset topspace_subtopology_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5078
    ultimately show "subtopology X S homeomorphic_space subtopology (prod_topology X euclideanreal) (f ` T)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5079
      by (smt (verit, best) closedin_subset f homT homeomorphic_map_subtopologies homeomorphic_space 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5080
          homeomorphic_space_trans topspace_subtopology topspace_subtopology_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5081
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5082
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5083
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5084
lemma completely_metrizable_space_gdelta_in_alt:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5085
  assumes X: "completely_metrizable_space X" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5086
    and S: "(countable intersection_of openin X) S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5087
  shows "completely_metrizable_space (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5088
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5089
  obtain \<U> where "countable \<U>" "S = \<Inter>\<U>" and ope: "\<And>U. U \<in> \<U> \<Longrightarrow> openin X U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5090
    using S by (force simp add: intersection_of_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5091
  then have \<U>: "completely_metrizable_space (powertop_real \<U>)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5092
    by (simp add: completely_metrizable_space_euclidean completely_metrizable_space_product_topology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5093
  obtain C where "closedin (prod_topology X (powertop_real \<U>)) C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5094
                and sub: "subtopology X (\<Inter>\<U>) homeomorphic_space
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5095
                   subtopology (prod_topology X (powertop_real \<U>)) C"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5096
    by (metis gdelta_homeomorphic_space_closedin_product  X completely_metrizable_imp_metrizable_space ope INF_identity_eq)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5097
  moreover have "completely_metrizable_space (prod_topology X (powertop_real \<U>))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5098
    by (simp add: completely_metrizable_space_prod_topology X \<U>)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5099
  ultimately have "completely_metrizable_space (subtopology (prod_topology X (powertop_real \<U>)) C)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5100
    using completely_metrizable_space_closedin by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5101
  then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5102
    using \<open>S = \<Inter>\<U>\<close> sub homeomorphic_completely_metrizable_space by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5103
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5104
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5105
lemma completely_metrizable_space_gdelta_in:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5106
   "\<lbrakk>completely_metrizable_space X; gdelta_in X S\<rbrakk>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5107
        \<Longrightarrow> completely_metrizable_space (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5108
  by (simp add: completely_metrizable_space_gdelta_in_alt gdelta_in_alt)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5109
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5110
lemma completely_metrizable_space_openin:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5111
   "\<lbrakk>completely_metrizable_space X; openin X S\<rbrakk>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5112
        \<Longrightarrow> completely_metrizable_space (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5113
  by (simp add: completely_metrizable_space_gdelta_in open_imp_gdelta_in)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5114
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5115
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5116
lemma (in Metric_space) locally_compact_imp_completely_metrizable_space:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5117
  assumes "locally_compact_space mtopology"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5118
  shows "completely_metrizable_space mtopology"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5119
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5120
  obtain f :: "['a,'a] \<Rightarrow> real" and m' where
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5121
    "mcomplete_of m'" and fim: "f \<in> M \<rightarrow> mspace m'"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5122
    and clo: "mtopology_of m' closure_of f ` M = mspace m'"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5123
    and d: "\<And>x y. \<lbrakk>x \<in> M; y \<in> M\<rbrakk> \<Longrightarrow> mdist m' (f x) (f y) = d x y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5124
    by (metis metric_completion)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5125
  then have "embedding_map mtopology (mtopology_of m') f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5126
    unfolding mtopology_of_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5127
    by (metis Metric_space12.isometry_imp_embedding_map Metric_space12_mspace_mdist mdist_metric mspace_metric)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5128
  then have hom: "mtopology homeomorphic_space subtopology (mtopology_of m') (f ` M)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5129
    by (metis embedding_map_imp_homeomorphic_space topspace_mtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5130
  have "locally_compact_space (subtopology (mtopology_of m') (f ` M))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5131
    using assms hom homeomorphic_locally_compact_space by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5132
  moreover have "Hausdorff_space (mtopology_of m')"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5133
    by (simp add: Metric_space.Hausdorff_space_mtopology mtopology_of_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5134
  ultimately have "openin (mtopology_of m') (f ` M)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5135
    by (simp add: clo dense_locally_compact_openin_Hausdorff_space fim image_subset_iff_funcset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5136
  then
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5137
  have "completely_metrizable_space (subtopology (mtopology_of m') (f ` M))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5138
    using \<open>mcomplete_of m'\<close> unfolding mcomplete_of_def mtopology_of_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5139
    by (metis Metric_space.completely_metrizable_space_mtopology Metric_space_mspace_mdist completely_metrizable_space_openin )
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5140
  then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5141
    using hom homeomorphic_completely_metrizable_space by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5142
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5143
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5144
lemma locally_compact_imp_completely_metrizable_space:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5145
  assumes "metrizable_space X" and "locally_compact_space X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5146
  shows "completely_metrizable_space X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5147
  by (metis Metric_space.locally_compact_imp_completely_metrizable_space assms metrizable_space_def)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5148
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5149
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5150
lemma completely_metrizable_space_imp_gdelta_in:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5151
  assumes X: "metrizable_space X" and "S \<subseteq> topspace X" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5152
    and XS: "completely_metrizable_space (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5153
  shows "gdelta_in X S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5154
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5155
  obtain U f where "gdelta_in X U" "S \<subseteq> U" and U: "U \<subseteq> X closure_of S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5156
            and contf: "continuous_map (subtopology X U) (subtopology X S) f" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5157
            and fid: "\<And>x. x \<in> S \<Longrightarrow> f x = x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5158
    using Lavrentiev_extension[of S X "subtopology X S" id] assms by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5159
  then have "f ` topspace (subtopology X U) \<subseteq> topspace (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5160
    using continuous_map_image_subset_topspace by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5161
  then have "f`U \<subseteq> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5162
    by (metis \<open>gdelta_in X U\<close> \<open>S \<subseteq> topspace X\<close> gdelta_in_subset topspace_subtopology_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5163
  moreover 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5164
  have "Hausdorff_space (subtopology X U)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5165
    by (simp add: Hausdorff_space_subtopology X metrizable_imp_Hausdorff_space)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5166
  then have "\<And>x. x \<in> U \<Longrightarrow> f x = x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5167
    using U fid contf forall_in_closure_of_eq [of _ "subtopology X U" S "subtopology X U" f id]
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5168
    by (metis \<open>S \<subseteq> U\<close> closure_of_subtopology_open continuous_map_id continuous_map_in_subtopology id_apply inf.orderE subtopology_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5169
  ultimately have "U \<subseteq> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5170
    by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5171
  then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5172
    using \<open>S \<subseteq> U\<close> \<open>gdelta_in X U\<close> by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5173
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5174
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5175
lemma completely_metrizable_space_eq_gdelta_in:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5176
   "\<lbrakk>completely_metrizable_space X; S \<subseteq> topspace X\<rbrakk>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5177
        \<Longrightarrow> completely_metrizable_space (subtopology X S) \<longleftrightarrow> gdelta_in X S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5178
  using completely_metrizable_imp_metrizable_space completely_metrizable_space_gdelta_in completely_metrizable_space_imp_gdelta_in by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5179
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5180
lemma gdelta_in_eq_completely_metrizable_space:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5181
   "completely_metrizable_space X
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5182
    \<Longrightarrow> gdelta_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> completely_metrizable_space (subtopology X S)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5183
  by (metis completely_metrizable_space_eq_gdelta_in gdelta_in_alt)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  5184
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5185
subsection\<open> Theorems from Kuratowski\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5186
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5187
text\<open>Kuratowski, Remark on an Invariance Theorem, \emph{Fundamenta Mathematicae} \textbf{37} (1950), pp. 251-252. 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5188
 The idea is that in suitable spaces, to show "number of components of the complement" (without 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5189
 distinguishing orders of infinity) is a homeomorphic invariant, it 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5190
 suffices to show it for closed subsets. Kuratowski states the main result 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5191
 for a "locally connected continuum", and seems clearly to be implicitly   
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5192
 assuming that means metrizable. We call out the general topological       
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5193
 hypotheses more explicitly, which do not however include connectedness.   \<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5194
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5195
lemma separation_by_closed_intermediates_count:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5196
  assumes X: "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5197
    and "finite \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5198
    and pwU: "pairwise (separatedin X) \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5199
    and nonempty: "{} \<notin> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5200
    and UU: "\<Union>\<U> = topspace X - S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5201
  obtains C where "closedin X C" "C \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5202
                  "\<And>D. \<lbrakk>closedin X D; C \<subseteq> D; D \<subseteq> S\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5203
                     \<Longrightarrow> \<exists>\<V>. \<V> \<approx> \<U> \<and> pairwise (separatedin X) \<V> \<and> {} \<notin> \<V> \<and> \<Union>\<V> = topspace X - D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5204
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5205
  obtain F where F: "\<And>S. S \<in> \<U> \<Longrightarrow> openin X (F S) \<and> S \<subseteq> F S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5206
    and pwF: "pairwise (\<lambda>S T. disjnt (F S) (F T)) \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5207
    using assms by (smt (verit, best) Diff_subset Sup_le_iff hereditarily_normal_separation_pairwise)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5208
  show thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5209
  proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5210
    show "closedin X (topspace X - \<Union>(F ` \<U>))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5211
      using F by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5212
    show "topspace X - \<Union>(F ` \<U>) \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5213
      using UU F by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5214
    show "\<exists>\<V>. \<V> \<approx> \<U> \<and> pairwise (separatedin X) \<V> \<and> {} \<notin> \<V> \<and> \<Union>\<V> = topspace X - C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5215
      if "closedin X C" "C \<subseteq> S" and C: "topspace X - \<Union>(F ` \<U>) \<subseteq> C" for C
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5216
    proof (intro exI conjI strip)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5217
      have "inj_on (\<lambda>S. F S - C) \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5218
        using pwF F
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5219
        unfolding inj_on_def pairwise_def disjnt_iff
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5220
        by (metis Diff_iff UU UnionI nonempty subset_empty subset_eq \<open>C \<subseteq> S\<close>)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5221
      then show "(\<lambda>S. F S - C) ` \<U> \<approx> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5222
        by simp
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5223
      show "pairwise (separatedin X) ((\<lambda>S. F S - C) ` \<U>)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5224
        using \<open>closedin X C\<close> F pwF by (force simp: pairwise_def openin_diff separatedin_open_sets disjnt_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5225
      show "{} \<notin> (\<lambda>S. F S - C) ` \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5226
        using nonempty UU \<open>C \<subseteq> S\<close> F
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5227
        by clarify (metis DiffD2 Diff_eq_empty_iff F UnionI subset_empty subset_eq)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5228
      show "(\<Union>S\<in>\<U>. F S - C) = topspace X - C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5229
        using UU F C openin_subset by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5230
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5231
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5232
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5233
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5234
lemma separation_by_closed_intermediates_gen:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5235
  assumes X: "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5236
    and discon: "\<not> connectedin X (topspace X - S)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5237
  obtains C where "closedin X C" "C \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5238
                  "\<And>D. \<lbrakk>closedin X D; C \<subseteq> D; D \<subseteq> S\<rbrakk> \<Longrightarrow> \<not> connectedin X (topspace X - D)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5239
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5240
  obtain C1 C2 where Ueq: "C1 \<union> C2 = topspace X - S" and "C1 \<noteq> {}" "C2 \<noteq> {}" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5241
    and sep: "separatedin X C1 C2" and "C1 \<noteq> C2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5242
    by (metis Diff_subset connectedin_eq_not_separated discon separatedin_refl)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5243
  then obtain C where "closedin X C" "C \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5244
    and C: "\<And>D. \<lbrakk>closedin X D; C \<subseteq> D; D \<subseteq> S\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5245
                     \<Longrightarrow> \<exists>\<V>. \<V> \<approx> {C1,C2} \<and> pairwise (separatedin X) \<V> \<and> {} \<notin> \<V> \<and> \<Union>\<V> = topspace X - D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5246
    using separation_by_closed_intermediates_count [of X "{C1,C2}" S] X
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5247
    apply (simp add: pairwise_insert separatedin_sym)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5248
    by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5249
  have "\<not> connectedin X (topspace X - D)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5250
    if D: "closedin X D" "C \<subseteq> D" "D \<subseteq> S" for D 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5251
  proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5252
    obtain V1 V2 where *: "pairwise (separatedin X) {V1,V2}" "{} \<notin> {V1,V2}" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5253
                          "\<Union>{V1,V2} = topspace X - D" "V1\<noteq>V2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5254
      by (metis C [OF D] \<open>C1 \<noteq> C2\<close> eqpoll_doubleton_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5255
    then have "disjnt V1 V2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5256
      by (metis pairwise_insert separatedin_imp_disjoint singleton_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5257
      with * show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5258
        by (auto simp add: connectedin_eq_not_separated pairwise_insert)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5259
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5260
  then show thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5261
    using \<open>C \<subseteq> S\<close> \<open>closedin X C\<close> that by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5262
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5263
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5264
lemma separation_by_closed_intermediates_eq_count:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5265
  fixes n::nat
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5266
  assumes lcX: "locally_connected_space X" and hnX: "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5267
  shows "(\<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - S) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5268
         (\<exists>C. closedin X C \<and> C \<subseteq> S \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5269
              (\<forall>D. closedin X D \<and> C \<subseteq> D \<and> D \<subseteq> S
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5270
                   \<longrightarrow> (\<exists>\<U>.  \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - D)))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5271
         (is "?lhs = ?rhs")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5272
proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5273
  assume ?lhs then show ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5274
    by (smt (verit, best) hnX separation_by_closed_intermediates_count eqpoll_iff_finite_card eqpoll_trans)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5275
next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5276
  assume R: ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5277
  show ?lhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5278
  proof (cases "n=0")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5279
    case True
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5280
    with R show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5281
      by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5282
  next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5283
    case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5284
    obtain C where "closedin X C" "C \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5285
             and C: "\<And>D. \<lbrakk>closedin X D; C \<subseteq> D; D \<subseteq> S\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5286
                      \<Longrightarrow> \<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5287
      using R by force
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5288
    then have "C \<subseteq> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5289
      by (simp add: closedin_subset)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5290
    define \<U> where "\<U> \<equiv> {D \<in> connected_components_of (subtopology X (topspace X - C)). D-S \<noteq> {}}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5291
    have ope\<U>: "openin X U" if "U \<in> \<U>" for U
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5292
      using that  \<open>closedin X C\<close> lcX locally_connected_space_open_connected_components 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5293
      by (fastforce simp add: closedin_def \<U>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5294
    have "{} \<notin> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5295
      by (auto simp: \<U>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5296
    have "pairwise disjnt \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5297
      using connected_components_of_disjoint by (fastforce simp add: pairwise_def \<U>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5298
    show ?lhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5299
    proof (rule ccontr)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5300
      assume con: "\<nexists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5301
      have card\<U>: "finite \<U> \<and> card \<U> < n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5302
      proof (rule ccontr)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5303
        assume "\<not> (finite \<U> \<and> card \<U> < n)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5304
        then obtain \<V> where "\<V> \<subseteq> \<U>" "finite \<V>" "card \<V> = n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5305
          by (metis infinite_arbitrarily_large linorder_not_less obtain_subset_with_card_n)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5306
        then obtain T where "T \<in> \<V>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5307
          using False by force
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5308
        define \<W> where "\<W> \<equiv> insert (topspace X - S - \<Union>(\<V> - {T})) ((\<lambda>D. D - S) ` (\<V> - {T}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5309
        have "\<Union>\<W> = topspace X - S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5310
          using \<open>\<And>U. U \<in> \<U> \<Longrightarrow> openin X U\<close> \<open>\<V> \<subseteq> \<U>\<close> topspace_def by (fastforce simp: \<W>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5311
        moreover have "{} \<notin> \<W>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5312
        proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5313
          obtain a where "a \<in> T" "a \<notin> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5314
            using \<U>_def \<open>T \<in> \<V>\<close> \<open>\<V> \<subseteq> \<U>\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5315
          then have "a \<in> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5316
            using \<open>T \<in> \<V>\<close> ope\<U> \<open>\<V> \<subseteq> \<U>\<close> openin_subset by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5317
          moreover have "a \<notin> \<Union>(\<V> - {T})"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5318
            using diff_Union_pairwise_disjoint [of \<V> "{T}"] \<open>disjoint \<U>\<close> pairwise_subset \<open>T \<in> \<V>\<close> \<open>\<V> \<subseteq> \<U>\<close> \<open>a \<in> T\<close> 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5319
            by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5320
          ultimately have "topspace X - S - \<Union>(\<V> - {T}) \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5321
            using \<open>a \<notin> S\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5322
          moreover have "\<And>V. V \<in> \<V> - {T} \<Longrightarrow> V - S \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5323
            using \<U>_def \<open>\<V> \<subseteq> \<U>\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5324
          ultimately show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5325
            by (metis (no_types, lifting) \<W>_def image_iff insert_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5326
        qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5327
        moreover have "disjoint \<V>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5328
          using \<open>\<V> \<subseteq> \<U>\<close> \<open>disjoint \<U>\<close> pairwise_subset by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5329
        then have inj: "inj_on (\<lambda>D. D - S) (\<V> - {T})"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5330
          unfolding inj_on_def using \<open>\<V> \<subseteq> \<U>\<close> disjointD \<U>_def inf_commute by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5331
        have "finite \<W>" "card \<W> = n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5332
          using \<open>{} \<notin> \<W>\<close> \<open>n \<noteq> 0\<close> \<open>T \<in> \<V>\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5333
          by (auto simp add: \<W>_def \<open>finite \<V>\<close> card_insert_if card_image inj \<open>card \<V> = n\<close>)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5334
        moreover have "pairwise (separatedin X) \<W>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5335
        proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5336
          have "disjoint \<W>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5337
            using \<open>disjoint \<V>\<close> by (auto simp: \<W>_def pairwise_def disjnt_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5338
          have "pairwise (separatedin (subtopology X (topspace X - S))) \<W>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5339
          proof (intro pairwiseI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5340
            fix A B
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5341
            assume \<section>: "A \<in> \<W>" "B \<in> \<W>" "A \<noteq> B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5342
            then have "disjnt A B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5343
              by (meson \<open>disjoint \<W>\<close> pairwiseD)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5344
            have "closedin (subtopology X (topspace X - C)) (\<Union>(\<V> - {T}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5345
              using \<U>_def \<open>\<V> \<subseteq> \<U>\<close> closedin_connected_components_of \<open>finite \<V>\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5346
              by (force simp add: intro!: closedin_Union)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5347
            with \<open>C \<subseteq> S\<close> have "openin (subtopology X (topspace X - S)) (topspace X - S - \<Union>(\<V> - {T}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5348
              by (fastforce simp add: openin_closedin_eq closedin_subtopology Int_absorb1)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5349
            moreover have "\<And>V. V \<in> \<V> \<and> V\<noteq>T \<Longrightarrow> openin (subtopology X (topspace X - S)) (V - S)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5350
              using \<open>\<V> \<subseteq> \<U>\<close> ope\<U>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5351
              by (metis IntD2 Int_Diff inf.orderE openin_subset openin_subtopology) 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5352
            ultimately have "openin (subtopology X (topspace X - S)) A" "openin (subtopology X (topspace X - S)) B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5353
              using \<section> \<W>_def by blast+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5354
            with \<open>disjnt A B\<close> show "separatedin (subtopology X (topspace X - S)) A B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5355
              using separatedin_open_sets by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5356
          qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5357
          then show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5358
            by (simp add: pairwise_def separatedin_subtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5359
        qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5360
        ultimately show False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5361
          by (metis con eqpoll_iff_finite_card)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5362
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5363
      obtain \<V> where "\<V> \<approx> {..<n} " "{} \<notin> \<V>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5364
                and pw\<V>: "pairwise (separatedin X) \<V>" and UV: "\<Union>\<V> = topspace X - (topspace X - \<Union>\<U>)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5365
      proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5366
        have "closedin X (topspace X - \<Union>\<U>)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5367
          using ope\<U> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5368
        moreover have "C \<subseteq> topspace X - \<Union>\<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5369
          using \<open>C \<subseteq> topspace X\<close> connected_components_of_subset by (fastforce simp: \<U>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5370
        moreover have "topspace X - \<Union>\<U> \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5371
          using Union_connected_components_of [of "subtopology X (topspace X - C)"] \<open>C \<subseteq> S\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5372
          by (auto simp: \<U>_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5373
        ultimately show thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5374
          by (metis C that)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5375
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5376
      have "\<V> \<lesssim> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5377
      proof (rule lepoll_relational_full)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5378
        have "\<Union>\<V> = \<Union>\<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5379
          by (simp add: Sup_le_iff UV double_diff ope\<U> openin_subset)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5380
        then show "\<exists>U. U \<in> \<U> \<and> \<not> disjnt U V" if "V \<in> \<V>" for V
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5381
          using that
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5382
          by (metis \<open>{} \<notin> \<V>\<close> disjnt_Union1 disjnt_self_iff_empty)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5383
        show "C1 = C2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5384
          if "T \<in> \<U>" and "C1 \<in> \<V>" and "C2 \<in> \<V>" and "\<not> disjnt T C1" and "\<not> disjnt T C2" for T C1 C2
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5385
        proof (cases "C1=C2")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5386
          case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5387
          then have "connectedin X T"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5388
            using \<U>_def connectedin_connected_components_of connectedin_subtopology \<open>T \<in> \<U>\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5389
          have "T \<subseteq> C1 \<union> \<Union>(\<V> - {C1})"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5390
            using \<open>\<Union>\<V> = \<Union>\<U>\<close> \<open>T \<in> \<U>\<close> by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5391
          with \<open>connectedin X T\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5392
          have "\<not> separatedin X C1 (\<Union>(\<V> - {C1}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5393
            unfolding connectedin_eq_not_separated_subset
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5394
            by (smt (verit) that False disjnt_def UnionI disjnt_iff insertE insert_Diff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5395
          with that show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5396
            by (metis (no_types, lifting) \<open>\<V> \<approx> {..<n}\<close> eqpoll_iff_finite_card finite_Diff pairwiseD pairwise_alt pw\<V> separatedin_Union(1) separatedin_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5397
        qed auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5398
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5399
      then show False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5400
        by (metis \<open>\<V> \<approx> {..<n}\<close> card\<U> eqpoll_iff_finite_card leD lepoll_iff_card_le)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5401
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5402
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5403
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5404
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5405
lemma separation_by_closed_intermediates_eq_gen:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5406
  assumes "locally_connected_space X" "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5407
  shows "\<not> connectedin X (topspace X - S) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5408
         (\<exists>C. closedin X C \<and> C \<subseteq> S \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5409
              (\<forall>D. closedin X D \<and> C \<subseteq> D \<and> D \<subseteq> S \<longrightarrow> \<not> connectedin X (topspace X - D)))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5410
    (is "?lhs = ?rhs")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5411
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5412
  have *: "(\<exists>\<U>::'a set set. \<U> \<approx> {..<Suc (Suc 0)} \<and> P \<U>) \<longleftrightarrow> (\<exists>A B. A\<noteq>B \<and> P{A,B})" for P
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5413
    by (metis One_nat_def eqpoll_doubleton_iff lessThan_Suc lessThan_empty_iff zero_neq_one)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5414
  have *: "(\<exists>C1 C2. separatedin X C1 C2 \<and> C1\<noteq>C2 \<and> C1\<noteq>{} \<and> C2\<noteq>{} \<and> C1 \<union> C2 = topspace X - S) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5415
         (\<exists>C. closedin X C \<and> C \<subseteq> S \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5416
              (\<forall>D. closedin X D \<and> C \<subseteq> D \<and> D \<subseteq> S
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5417
                   \<longrightarrow> (\<exists>C1 C2. separatedin X C1 C2 \<and> C1\<noteq>C2 \<and> C1\<noteq>{} \<and> C2\<noteq>{} \<and> C1 \<union> C2 = topspace X - D)))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5418
    using separation_by_closed_intermediates_eq_count [OF assms, of "Suc(Suc 0)" S]
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5419
    apply (simp add: * pairwise_insert separatedin_sym cong: conj_cong)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5420
    apply (simp add: eq_sym_conv conj_ac)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5421
    done
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5422
  with separatedin_refl
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5423
  show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5424
    apply (simp add: connectedin_eq_not_separated)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5425
    by (smt (verit, best) separatedin_refl)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5426
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5427
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5428
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5429
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5430
lemma lepoll_connnected_components_connectedin:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5431
  assumes "\<And>C. C \<in> \<U> \<Longrightarrow> connectedin X C"  "\<Union>\<U> = topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5432
  shows "connected_components_of X \<lesssim> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5433
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5434
  have "connected_components_of X \<lesssim> \<U> - {{}}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5435
  proof (rule lepoll_relational_full)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5436
    show "\<exists>U. U \<in> \<U> - {{}} \<and> U \<subseteq> V"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5437
      if "V \<in> connected_components_of X" for V 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5438
      using that unfolding connected_components_of_def image_iff
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5439
      by (metis Union_iff assms connected_component_of_maximal empty_iff insert_Diff_single insert_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5440
    show "V = V'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5441
      if "U \<in> \<U> - {{}}" "V \<in> connected_components_of X" "V' \<in> connected_components_of X" "U \<subseteq> V" "U \<subseteq> V'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5442
      for U V V'
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5443
      by (metis DiffD2 disjointD insertCI le_inf_iff pairwise_disjoint_connected_components_of subset_empty that)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5444
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5445
  also have "\<dots> \<lesssim> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5446
    by (simp add: subset_imp_lepoll)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5447
  finally show ?thesis .
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5448
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5449
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5450
lemma lepoll_connected_components_alt:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5451
  "{..<n::nat} \<lesssim> connected_components_of X \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5452
    n = 0 \<or> (\<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5453
  (is "?lhs \<longleftrightarrow> ?rhs")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5454
proof (cases "n=0")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5455
next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5456
  case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5457
  show ?thesis 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5458
  proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5459
    assume L: ?lhs 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5460
    with False show ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5461
    proof (induction n rule: less_induct)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5462
      case (less n)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5463
      show ?case
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5464
      proof (cases "n\<le>1")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5465
        case True
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5466
        with less.prems have "topspace X \<noteq> {}" "n=1"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5467
          by (fastforce simp add: connected_components_of_def)+
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5468
        then have "{} \<notin> {topspace X}"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5469
          by blast
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5470
        with \<open>n=1\<close> show ?thesis
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
  5471
          by (simp add: eqpoll_iff_finite_card card_Suc_eq flip: ex_simps)
78274
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5472
      next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5473
        case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5474
        then have "n-1 \<noteq> 0"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5475
          by linarith
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5476
        have n1_lesspoll: "{..<n-1} \<prec> {..<n}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5477
          using False lesspoll_iff_finite_card by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5478
        also have "\<dots> \<lesssim> connected_components_of X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5479
          using less by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5480
        finally have "{..<n-1} \<lesssim> connected_components_of X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5481
          using lesspoll_imp_lepoll by blast 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5482
        then obtain \<U> where Ueq: "\<U> \<approx> {..<n-1}" and "{} \<notin> \<U>" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5483
          and pwU: "pairwise (separatedin X) \<U>" and UU: "\<Union>\<U> = topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5484
          by (meson \<open>n - 1 \<noteq> 0\<close> diff_less gr0I less zero_less_one)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5485
        show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5486
        proof (cases "\<forall>C \<in> \<U>. connectedin X C")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5487
          case True
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5488
          then show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5489
            using lepoll_connnected_components_connectedin [of \<U> X] less.prems
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5490
            by (metis UU Ueq lepoll_antisym lepoll_trans lepoll_trans2 lesspoll_def n1_lesspoll)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5491
          next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5492
            case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5493
            with UU obtain C A B where ABC: "C \<in> \<U>" "A \<union> B = C" "A \<noteq> {}" "B \<noteq> {}" and sep: "separatedin X A B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5494
              by (fastforce simp add: connectedin_eq_not_separated)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5495
            define \<V> where "\<V> \<equiv> insert A (insert B (\<U> - {C}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5496
            have "\<V> \<approx> {..<n}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5497
            proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5498
              have "A \<noteq> B"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5499
                using \<open>B \<noteq> {}\<close> sep by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5500
              moreover obtain "A \<notin> \<U>" "B \<notin> \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5501
                using pwU unfolding pairwise_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5502
                by (metis ABC sep separatedin_Un(1) separatedin_refl separatedin_sym)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5503
              moreover have "card \<U> = n-1" "finite \<U>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5504
                using Ueq eqpoll_iff_finite_card by blast+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5505
              ultimately
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5506
              have "card (insert A (insert B (\<U> - {C}))) = n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5507
                using \<open>C \<in> \<U>\<close> by (auto simp add: card_insert_if)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5508
              then show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5509
                using \<V>_def \<open>finite \<U>\<close> eqpoll_iff_finite_card by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5510
            qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5511
            moreover have "{} \<notin> \<V>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5512
              using ABC \<V>_def \<open>{} \<notin> \<U>\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5513
            moreover have "\<Union>\<V> = topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5514
              using ABC UU \<V>_def by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5515
            moreover have "pairwise (separatedin X) \<V>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5516
              using pwU sep ABC unfolding  \<V>_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5517
              apply (simp add: separatedin_sym pairwise_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5518
              by (metis member_remove remove_def separatedin_Un(1))
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5519
            ultimately show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5520
              by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5521
          qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5522
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5523
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5524
  next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5525
    assume ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5526
    then obtain \<U> where "\<U> \<approx> {..<n}" "{} \<notin> \<U>" and pwU: "pairwise (separatedin X) \<U>" and UU: "\<Union>\<U> = topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5527
      using False by force
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5528
    have "card (connected_components_of X) \<ge> n" if "finite (connected_components_of X)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5529
    proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5530
      have "\<U> \<lesssim> connected_components_of X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5531
      proof (rule lepoll_relational_full)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5532
        show "\<exists>T. T \<in> connected_components_of X \<and> \<not> disjnt T C" if "C \<in> \<U>" for C 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5533
          by (metis that UU Union_connected_components_of Union_iff \<open>{} \<notin> \<U>\<close> disjnt_iff equals0I)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5534
        show "(C1::'a set) = C2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5535
          if "T \<in> connected_components_of X" and "C1 \<in> \<U>" "C2 \<in> \<U>" "\<not> disjnt T C1" "\<not> disjnt T C2" for T C1 C2
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5536
        proof (rule ccontr)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5537
          assume "C1 \<noteq> C2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5538
          then have "connectedin X T"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5539
            by (simp add: connectedin_connected_components_of that(1))
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5540
          moreover have "\<not> separatedin X C1 (\<Union>(\<U> - {C1}))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5541
            using \<open>connectedin X T\<close> pwU unfolding pairwise_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5542
            by (smt (verit) Sup_upper UU Union_connected_components_of \<open>C1 \<noteq> C2\<close> complete_lattice_class.Sup_insert connectedin_subset_separated_union disjnt_subset2 disjnt_sym insert_Diff separatedin_imp_disjoint that)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5543
          ultimately show False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5544
            using \<open>\<U> \<approx> {..<n}\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5545
            apply (simp add: connectedin_eq_not_separated_subset eqpoll_iff_finite_card)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5546
            by (metis Sup_upper UU finite_Diff pairwise_alt pwU separatedin_Union(1) that(2))
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5547
        qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5548
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5549
      then show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5550
        by (metis \<open>\<U> \<approx> {..<n}\<close> eqpoll_iff_finite_card lepoll_iff_card_le that)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5551
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5552
    then show ?lhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5553
      by (metis card_lessThan finite_lepoll_infinite finite_lessThan lepoll_iff_card_le)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5554
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5555
qed auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5556
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5557
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5558
subsection\<open>A perfect set in common cases must have at least the cardinality of the continuum\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5559
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5560
lemma (in Metric_space) lepoll_perfect_set:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5561
  assumes "mcomplete"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5562
    and "mtopology derived_set_of S = S" "S \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5563
  shows "(UNIV::real set) \<lesssim> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5564
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5565
  have "S \<subseteq> M"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5566
    using assms(2) derived_set_of_infinite_mball by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5567
  have "(UNIV::real set) \<lesssim> (UNIV::nat set set)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5568
    using eqpoll_imp_lepoll eqpoll_sym nat_sets_eqpoll_reals by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5569
  also have "\<dots> \<lesssim> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5570
  proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5571
    have "\<exists>y z \<delta>. y \<in> S \<and> z \<in> S \<and> 0 < \<delta> \<and> \<delta> < \<epsilon>/2 \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5572
                  mcball y \<delta> \<subseteq> mcball x \<epsilon> \<and> mcball z \<delta> \<subseteq> mcball x \<epsilon> \<and> disjnt (mcball y \<delta>) (mcball z \<delta>)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5573
      if "x \<in> S" "0 < \<epsilon>" for x \<epsilon>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5574
    proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5575
      define S' where "S' \<equiv> S \<inter> mball x (\<epsilon>/4)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5576
      have "infinite S'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5577
        using derived_set_of_infinite_mball [of S] assms that S'_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5578
        by (smt (verit, ccfv_SIG) mem_Collect_eq zero_less_divide_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5579
      then have "\<And>x y z. \<not> (S' \<subseteq> {x,y,z})"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5580
        using finite_subset by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5581
      then obtain l r where lr: "l \<in> S'" "r \<in> S'" "l\<noteq>r" "l\<noteq>x" "r\<noteq>x"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5582
        by (metis insert_iff subsetI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5583
      show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5584
      proof (intro exI conjI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5585
        show "l \<in> S" "r \<in> S" "d l r / 3 > 0"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5586
          using lr by (auto simp: S'_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5587
        show "d l r / 3 < \<epsilon>/2" "mcball l (d l r / 3) \<subseteq> mcball x \<epsilon>" "mcball r (d l r / 3) \<subseteq> mcball x \<epsilon>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5588
          using lr by (clarsimp simp: S'_def, smt (verit) commute triangle'')+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5589
        show "disjnt (mcball l (d l r / 3)) (mcball r (d l r / 3))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5590
          using lr by (simp add: S'_def disjnt_iff) (smt (verit, best) mdist_pos_less triangle')
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5591
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5592
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5593
    then obtain l r \<delta> 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5594
        where lrS: "\<And>x \<epsilon>. \<lbrakk>x \<in> S; 0 < \<epsilon>\<rbrakk> \<Longrightarrow> l x \<epsilon> \<in> S \<and> r x \<epsilon> \<in> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5595
          and \<delta>: "\<And>x \<epsilon>. \<lbrakk>x \<in> S; 0 < \<epsilon>\<rbrakk> \<Longrightarrow> 0 < \<delta> x \<epsilon> \<and> \<delta> x \<epsilon> < \<epsilon>/2"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5596
          and "\<And>x \<epsilon>. \<lbrakk>x \<in> S; 0 < \<epsilon>\<rbrakk> \<Longrightarrow>  mcball (l x \<epsilon>) (\<delta> x \<epsilon>) \<subseteq> mcball x \<epsilon> \<and> mcball (r x \<epsilon>) (\<delta> x \<epsilon>) \<subseteq> mcball x \<epsilon> \<and> 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5597
                  disjnt (mcball (l x \<epsilon>) (\<delta> x \<epsilon>)) (mcball (r x \<epsilon>) (\<delta> x \<epsilon>))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5598
      by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5599
    then have lr_mcball: "\<And>x \<epsilon>. \<lbrakk>x \<in> S; 0 < \<epsilon>\<rbrakk> \<Longrightarrow> mcball (l x \<epsilon>) (\<delta> x \<epsilon>) \<subseteq> mcball x \<epsilon> \<and> mcball (r x \<epsilon>) (\<delta> x \<epsilon>) \<subseteq> mcball x \<epsilon> "
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5600
          and lr_disjnt: "\<And>x \<epsilon>. \<lbrakk>x \<in> S; 0 < \<epsilon>\<rbrakk> \<Longrightarrow> disjnt (mcball (l x \<epsilon>) (\<delta> x \<epsilon>)) (mcball (r x \<epsilon>) (\<delta> x \<epsilon>))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5601
      by metis+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5602
    obtain a where "a \<in> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5603
      using \<open>S \<noteq> {}\<close> by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5604
    define xe where "xe \<equiv> 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5605
           \<lambda>B. rec_nat (a,1) (\<lambda>n (x,\<gamma>). ((if n\<in>B then r else l) x \<gamma>, \<delta> x \<gamma>))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5606
    have [simp]: "xe b 0 = (a,1)" for b
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5607
      by (simp add: xe_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5608
    have "xe B (Suc n) = (let (x,\<gamma>) = xe B n in ((if n\<in>B then r else l) x \<gamma>, \<delta> x \<gamma>))" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5609
      by (simp add: xe_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5610
    define x where "x \<equiv> \<lambda>B n. fst (xe B n)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5611
    define \<gamma> where "\<gamma> \<equiv> \<lambda>B n. snd (xe B n)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5612
    have [simp]: "x B 0 = a" "\<gamma> B 0 = 1" for B
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5613
      by (simp_all add: x_def \<gamma>_def xe_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5614
    have x_Suc[simp]: "x B (Suc n) = ((if n\<in>B then r else l) (x B n) (\<gamma> B n))" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5615
     and \<gamma>_Suc[simp]: "\<gamma> B (Suc n) = \<delta> (x B n) (\<gamma> B n)" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5616
      by (simp_all add: x_def \<gamma>_def xe_def split: prod.split)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5617
    interpret Submetric M d S
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5618
    proof qed (use \<open>S \<subseteq> M\<close> in metis)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5619
    have "closedin mtopology S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5620
      by (metis assms(2) closure_of closure_of_eq inf.absorb_iff2 subset subset_Un_eq subset_refl topspace_mtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5621
    with \<open>mcomplete\<close>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5622
    have "sub.mcomplete"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5623
      by (metis closedin_mcomplete_imp_mcomplete)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5624
    have *: "x B n \<in> S \<and> \<gamma> B n > 0" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5625
      by (induction n) (auto simp: \<open>a \<in> S\<close> lrS \<delta>)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5626
    with subset have E: "x B n \<in> M" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5627
      by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5628
    have \<gamma>_le: "\<gamma> B n \<le> (1/2)^n" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5629
    proof(induction n)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5630
      case 0 then show ?case by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5631
    next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5632
      case (Suc n)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5633
      then show ?case
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5634
        by simp (smt (verit) "*" \<delta> field_sum_of_halves)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5635
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5636
    { fix B
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5637
      have "\<And>n. sub.mcball (x B (Suc n)) (\<gamma> B (Suc n)) \<subseteq> sub.mcball (x B n) (\<gamma> B n)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5638
        by (smt (verit, best) "*" Int_iff \<gamma>_Suc x_Suc in_mono lr_mcball mcball_submetric_eq subsetI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5639
      then have mon: "monotone (\<le>) (\<lambda>x y. y \<subseteq> x) (\<lambda>n. sub.mcball (x B n) (\<gamma> B n))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5640
        by (simp add: decseq_SucI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5641
      have "\<exists>n a. sub.mcball (x B n) (\<gamma> B n) \<subseteq> sub.mcball a \<epsilon>" if "\<epsilon>>0" for \<epsilon>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5642
      proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5643
        obtain n where "(1/2)^n < \<epsilon>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5644
          using \<open>0 < \<epsilon>\<close> real_arch_pow_inv by force
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5645
        with \<gamma>_le have \<epsilon>: "\<gamma> B n \<le> \<epsilon>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5646
          by (smt (verit))
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5647
        show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5648
        proof (intro exI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5649
          show "sub.mcball (x B n) (\<gamma> B n) \<subseteq> sub.mcball (x B n) \<epsilon>"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5650
            by (simp add: \<epsilon> sub.mcball_subset_concentric)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5651
        qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5652
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5653
      then have "\<exists>l. l \<in> S \<and> (\<Inter>n. sub.mcball (x B n) (\<gamma> B n)) = {l}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5654
        using \<open>sub.mcomplete\<close> mon 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5655
        unfolding sub.mcomplete_nest_sing
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5656
        apply (drule_tac x="\<lambda>n. sub.mcball (x B n) (\<gamma> B n)" in spec)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5657
        by (meson * order.asym sub.closedin_mcball sub.mcball_eq_empty)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5658
    }
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5659
    then obtain z where z: "\<And>B. z B \<in> S \<and> (\<Inter>n. sub.mcball (x B n) (\<gamma> B n)) = {z B}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5660
      by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5661
    show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5662
      unfolding lepoll_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5663
    proof (intro exI conjI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5664
      show "inj z"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5665
      proof (rule inj_onCI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5666
        fix B C
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5667
        assume eq: "z B = z C" and "B \<noteq> C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5668
        then have ne: "sym_diff B C \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5669
          by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5670
        define n where "n \<equiv> LEAST k. k \<in> (sym_diff B C)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5671
        with ne have n: "n \<in> sym_diff B C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5672
          by (metis Inf_nat_def1 LeastI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5673
        then have non: "n \<in> B \<longleftrightarrow> n \<notin> C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5674
          by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5675
        have H: "z C \<in> sub.mcball (x B (Suc n)) (\<gamma> B (Suc n)) \<and> z C \<in> sub.mcball (x C (Suc n)) (\<gamma> C (Suc n))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5676
          using z [of B] z [of C] apply (simp add: lrS set_eq_iff non *)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5677
          by (smt (verit, best) \<gamma>_Suc eq non x_Suc)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5678
        have "k \<in> B \<longleftrightarrow> k \<in> C" if "k<n" for k
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5679
          using that unfolding n_def by (meson DiffI UnCI not_less_Least)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5680
        moreover have "(\<forall>m. m < p \<longrightarrow> (m \<in> B \<longleftrightarrow> m \<in> C)) \<Longrightarrow> x B p = x C p \<and> \<gamma> B p = \<gamma> C p" for p
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5681
          by (induction p) auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5682
        ultimately have "x B n = x C n" "\<gamma> B n = \<gamma> C n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5683
           by blast+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5684
        then show False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5685
          using lr_disjnt * H non
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5686
          by (smt (verit) IntD2 \<gamma>_Suc disjnt_iff mcball_submetric_eq x_Suc)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5687
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5688
      show "range z \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5689
        using z by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5690
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5691
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5692
  finally show ?thesis .
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5693
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5694
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5695
lemma lepoll_perfect_set_aux:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5696
  assumes lcX: "locally_compact_space X" and hsX: "Hausdorff_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5697
    and eq: "X derived_set_of topspace X = topspace X" and "topspace X \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5698
  shows "(UNIV::real set) \<lesssim> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5699
proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5700
  have "(UNIV::real set) \<lesssim> (UNIV::nat set set)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5701
    using eqpoll_imp_lepoll eqpoll_sym nat_sets_eqpoll_reals by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5702
  also have "\<dots> \<lesssim> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5703
  proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5704
    obtain z where z: "z \<in> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5705
      using assms by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5706
    then obtain U K where "openin X U" "compactin X K" "U \<noteq> {}" "U \<subseteq> K"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5707
      by (metis emptyE lcX locally_compact_space_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5708
    then have "closedin X K"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5709
      by (simp add: compactin_imp_closedin hsX)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5710
    have intK_ne: "X interior_of K \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5711
        using \<open>U \<noteq> {}\<close> \<open>U \<subseteq> K\<close> \<open>openin X U\<close> interior_of_eq_empty by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5712
    have "\<exists>D E. closedin X D \<and> D \<subseteq> K \<and> X interior_of D \<noteq> {} \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5713
                closedin X E \<and> E \<subseteq> K \<and> X interior_of E \<noteq> {} \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5714
                disjnt D E \<and> D \<subseteq> C \<and> E \<subseteq> C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5715
      if "closedin X C" "C \<subseteq> K" and C: "X interior_of C \<noteq> {}" for C
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5716
    proof -
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5717
      obtain z where z: "z \<in> X interior_of C" "z \<in> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5718
        using C interior_of_subset_topspace by fastforce 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5719
      obtain x y where "x \<in> X interior_of C" "y \<in> X interior_of C" "x\<noteq>y"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5720
        by (metis z eq in_derived_set_of openin_interior_of)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5721
      then have "x \<in> topspace X" "y \<in> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5722
        using interior_of_subset_topspace by force+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5723
      with hsX obtain V W where "openin X V" "openin X W" "x \<in> V" "y \<in> W" "disjnt V W"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5724
        by (metis Hausdorff_space_def \<open>x \<noteq> y\<close>)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5725
      have *: "\<And>W x. openin X W \<and> x \<in> W
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5726
            \<Longrightarrow> \<exists>U V. openin X U \<and> closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5727
        using lcX hsX locally_compact_Hausdorff_imp_regular_space neighbourhood_base_of_closedin neighbourhood_base_of
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5728
        by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5729
      obtain M D where MD: "openin X M" "closedin X D" "y \<in> M" "M \<subseteq> D" "D \<subseteq> X interior_of C \<inter> W"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5730
        using * [of "X interior_of C \<inter> W" y]
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5731
        using \<open>openin X W\<close> \<open>y \<in> W\<close> \<open>y \<in> X interior_of C\<close> by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5732
      obtain N E where NE: "openin X N" "closedin X E" "x \<in> N" "N \<subseteq> E" "E \<subseteq> X interior_of C \<inter> V"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5733
        using * [of "X interior_of C \<inter> V" x]
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5734
        using \<open>openin X V\<close> \<open>x \<in> V\<close> \<open>x \<in> X interior_of C\<close> by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5735
      show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5736
      proof (intro exI conjI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5737
        show "X interior_of D \<noteq> {}" "X interior_of E \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5738
          using MD NE by (fastforce simp: interior_of_def)+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5739
        show "disjnt D E"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5740
          by (meson MD(5) NE(5) \<open>disjnt V W\<close> disjnt_subset1 disjnt_sym le_inf_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5741
      qed (use MD NE \<open>C \<subseteq> K\<close> interior_of_subset in force)+
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5742
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5743
    then obtain L R where
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5744
     LR: "\<And>C. \<lbrakk>closedin X C; C \<subseteq> K; X interior_of C \<noteq> {}\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5745
      \<Longrightarrow> closedin X (L C) \<and> (L C) \<subseteq> K \<and> X interior_of (L C) \<noteq> {} \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5746
                closedin X (R C) \<and> (R C) \<subseteq> K \<and> X interior_of (R C) \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5747
     and disjLR: "\<And>C. \<lbrakk>closedin X C; C \<subseteq> K; X interior_of C \<noteq> {}\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5748
      \<Longrightarrow> disjnt (L C) (R C) \<and> (L C) \<subseteq> C \<and> (R C) \<subseteq> C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5749
      by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5750
    define d where "d \<equiv> \<lambda>B. rec_nat K (\<lambda>n. if n \<in> B then R else L)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5751
    have d0[simp]: "d B 0 = K" for B
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5752
      by (simp add: d_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5753
    have [simp]: "d B (Suc n) = (if n \<in> B then R else L) (d B n)" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5754
      by (simp add: d_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5755
    have d_correct: "closedin X (d B n) \<and> d B n \<subseteq> K \<and> X interior_of (d B n) \<noteq> {}" for B n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5756
    proof (induction n)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5757
      case 0
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5758
      then show ?case by (auto simp: \<open>closedin X K\<close> intK_ne)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5759
    next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5760
      case (Suc n) with LR show ?case by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5761
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5762
    have "(\<Inter>n. d B n) \<noteq> {}" for B
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5763
    proof (rule compact_space_imp_nest)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5764
      show "compact_space (subtopology X K)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5765
        by (simp add: \<open>compactin X K\<close> compact_space_subtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5766
      show "closedin (subtopology X K) (d B n)" for n :: nat
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5767
        by (simp add: closedin_subset_topspace d_correct)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5768
      show "d B n \<noteq> {}" for n :: nat
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5769
        by (metis d_correct interior_of_empty)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5770
      show "antimono (d B)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5771
      proof (rule antimonoI [OF transitive_stepwise_le])
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5772
        fix n
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5773
        show "d B (Suc n) \<subseteq> d B n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5774
        by (simp add: d_correct disjLR)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5775
      qed auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5776
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5777
    then obtain x where x: "\<And>B. x B \<in> (\<Inter>n. d B n)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5778
      unfolding set_eq_iff by (metis empty_iff)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5779
    show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5780
      unfolding lepoll_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5781
    proof (intro exI conjI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5782
      show "inj x"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5783
      proof (rule inj_onCI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5784
        fix B C
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5785
        assume eq: "x B = x C" and "B\<noteq>C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5786
        then have ne: "sym_diff B C \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5787
          by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5788
        define n where "n \<equiv> LEAST k. k \<in> (sym_diff B C)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5789
        with ne have n: "n \<in> sym_diff B C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5790
          by (metis Inf_nat_def1 LeastI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5791
        then have non: "n \<in> B \<longleftrightarrow> n \<notin> C"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5792
          by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5793
        have "k \<in> B \<longleftrightarrow> k \<in> C" if "k<n" for k
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5794
          using that unfolding n_def by (meson DiffI UnCI not_less_Least)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5795
        moreover have "(\<forall>m. m < p \<longrightarrow> (m \<in> B \<longleftrightarrow> m \<in> C)) \<Longrightarrow> d B p = d C p" for p
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5796
          by (induction p) auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5797
        ultimately have "d B n = d C n"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5798
          by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5799
        then have "disjnt (d B (Suc n)) (d C (Suc n))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5800
          by (simp add: d_correct disjLR disjnt_sym non)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5801
        then show False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5802
          by (metis InterE disjnt_iff eq rangeI x)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5803
      qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5804
      show "range x \<subseteq> topspace X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5805
        using x d0 \<open>compactin X K\<close> compactin_subset_topspace d_correct by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5806
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5807
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5808
  finally show ?thesis .
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5809
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5810
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5811
lemma lepoll_perfect_set:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5812
  assumes X: "completely_metrizable_space X \<or> locally_compact_space X \<and> Hausdorff_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5813
    and S: "X derived_set_of S = S" "S \<noteq> {}"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5814
  shows "(UNIV::real set) \<lesssim> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5815
  using X
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5816
proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5817
  assume "completely_metrizable_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5818
  with assms show "(UNIV::real set) \<lesssim> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5819
    by (metis Metric_space.lepoll_perfect_set completely_metrizable_space_def)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5820
next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5821
  assume "locally_compact_space X \<and> Hausdorff_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5822
  then show "(UNIV::real set) \<lesssim> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5823
    using lepoll_perfect_set_aux [of "subtopology X S"]
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5824
    by (metis Hausdorff_space_subtopology S closedin_derived_set_of closedin_subset derived_set_of_subtopology 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5825
        locally_compact_space_closed_subset subtopology_topspace topspace_subtopology topspace_subtopology_subset)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5826
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5827
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5828
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5829
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5830
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5831
lemma Kuratowski_aux1:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5832
  assumes "\<And>S T. R S T \<Longrightarrow> R T S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5833
  shows "(\<forall>S T n. R S T \<longrightarrow> (f S \<approx> {..<n::nat} \<longleftrightarrow> f T \<approx> {..<n::nat})) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5834
         (\<forall>n S T. R S T \<longrightarrow> {..<n::nat} \<lesssim> f S \<longrightarrow> {..<n::nat} \<lesssim> f T)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5835
         (is "?lhs = ?rhs")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5836
proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5837
  assume ?lhs then show ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5838
    by (meson eqpoll_iff_finite_card eqpoll_sym finite_lepoll_infinite finite_lessThan lepoll_trans2)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5839
next
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5840
  assume ?rhs then show ?lhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5841
    by (smt (verit, best) lepoll_iff_finite_card  assms eqpoll_iff_finite_card finite_lepoll_infinite 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5842
        finite_lessThan le_Suc_eq lepoll_antisym lepoll_iff_card_le not_less_eq_eq)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5843
qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5844
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5845
lemma Kuratowski_aux2:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5846
   "pairwise (separatedin (subtopology X (topspace X - S))) \<U> \<and> {} \<notin> \<U> \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5847
     \<Union>\<U> = topspace(subtopology X (topspace X - S)) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5848
     pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5849
  by (auto simp: pairwise_def separatedin_subtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5850
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5851
proposition Kuratowski_component_number_invariance_aux:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5852
  assumes "compact_space X" and HsX: "Hausdorff_space X" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5853
    and lcX: "locally_connected_space X" and hnX: "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5854
    and hom: "(subtopology X S) homeomorphic_space (subtopology X T)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5855
    and leXS: "{..<n::nat} \<lesssim> connected_components_of (subtopology X (topspace X - S))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5856
  assumes \<section>: "\<And>S T.
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5857
              \<lbrakk>closedin X S; closedin X T; (subtopology X S) homeomorphic_space (subtopology X T);
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5858
              {..<n::nat} \<lesssim> connected_components_of (subtopology X (topspace X - S))\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5859
              \<Longrightarrow> {..<n::nat} \<lesssim> connected_components_of (subtopology X (topspace X - T))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5860
  shows "{..<n::nat} \<lesssim> connected_components_of (subtopology X (topspace X - T))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5861
proof (cases "n=0")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5862
  case False
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5863
  obtain f g where homf: "homeomorphic_map (subtopology X S) (subtopology X T) f" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5864
      and homg: "homeomorphic_map (subtopology X T) (subtopology X S) g"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5865
    and gf: "\<And>x. x \<in> topspace (subtopology X S) \<Longrightarrow> g(f x) = x" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5866
    and fg: "\<And>y. y \<in> topspace (subtopology X T) \<Longrightarrow> f(g y) = y"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5867
    and f: "f \<in> topspace (subtopology X S) \<rightarrow> topspace (subtopology X T)" 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5868
    and g: "g \<in> topspace (subtopology X T) \<rightarrow> topspace (subtopology X S)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5869
    using homeomorphic_space_unfold hom by metis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5870
  obtain C where "closedin X C" "C \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5871
     and C: "\<And>D. \<lbrakk>closedin X D; C \<subseteq> D; D \<subseteq> S\<rbrakk>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5872
           \<Longrightarrow> \<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5873
    using separation_by_closed_intermediates_eq_count [of X n S] assms
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5874
    by (smt (verit, ccfv_threshold) False Kuratowski_aux2 lepoll_connected_components_alt)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5875
  have "\<exists>C. closedin X C \<and> C \<subseteq> T \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5876
          (\<forall>D. closedin X D \<and> C \<subseteq> D \<and> D \<subseteq> T
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5877
               \<longrightarrow> (\<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5878
                        {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - D))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5879
  proof (intro exI, intro conjI strip)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5880
    have "compactin X (f ` C)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5881
      by (meson \<open>C \<subseteq> S\<close> \<open>closedin X C\<close> assms(1) closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homf)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5882
    then show "closedin X (f ` C)"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5883
      using \<open>Hausdorff_space X\<close> compactin_imp_closedin by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5884
    show "f ` C \<subseteq> T"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5885
      by (meson \<open>C \<subseteq> S\<close> \<open>closedin X C\<close> closedin_imp_subset closedin_subset_topspace homeomorphic_map_closedness_eq homf)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5886
    fix D'
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5887
    assume D': "closedin X D' \<and> f ` C \<subseteq> D' \<and> D' \<subseteq> T"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5888
    define D where "D \<equiv> g ` D'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5889
    have "compactin X D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5890
      unfolding D_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5891
      by (meson D' \<open>compact_space X\<close> closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homg)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5892
    then have "closedin X D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5893
      by (simp add: assms(2) compactin_imp_closedin)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5894
    moreover have "C \<subseteq> D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5895
      using D' D_def \<open>C \<subseteq> S\<close> \<open>closedin X C\<close> closedin_subset gf image_iff by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5896
    moreover have "D \<subseteq> S"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5897
      by (metis D' D_def assms(1) closedin_compact_space compactin_subtopology homeomorphic_map_compactness_eq homg)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5898
    ultimately obtain \<U> where "\<U> \<approx> {..<n}" "pairwise (separatedin X) \<U>" "{} \<notin> \<U>" "\<Union>\<U> = topspace X - D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5899
      using C by meson
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5900
    moreover have "(subtopology X D) homeomorphic_space (subtopology X D')"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5901
      unfolding homeomorphic_space_def
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5902
    proof (intro exI)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5903
      have "subtopology X D = subtopology (subtopology X S) D"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5904
        by (simp add: \<open>D \<subseteq> S\<close> inf.absorb2 subtopology_subtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5905
      moreover have "subtopology X D' = subtopology (subtopology X T) D'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5906
        by (simp add: D' inf.absorb2 subtopology_subtopology)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5907
      moreover have "homeomorphic_maps (subtopology X T) (subtopology X S) g f"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5908
        by (simp add: fg gf homeomorphic_maps_map homf homg)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5909
      ultimately
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5910
      have "homeomorphic_maps (subtopology X D') (subtopology X D) g f"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5911
        by (metis D' D_def \<open>closedin X D\<close> closedin_subset homeomorphic_maps_subtopologies topspace_subtopology Int_absorb1)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5912
      then show "homeomorphic_maps (subtopology X D) (subtopology X D') f g"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5913
        using homeomorphic_maps_sym by blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5914
    qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5915
    ultimately show "\<exists>\<U>. \<U> \<approx> {..<n} \<and> pairwise (separatedin X) \<U> \<and> {} \<notin> \<U> \<and> \<Union> \<U> = topspace X - D'"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5916
      by (smt (verit, ccfv_SIG) \<section> D' False \<open>closedin X D\<close> Kuratowski_aux2 lepoll_connected_components_alt)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5917
  qed
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5918
  then have "\<exists>\<U>. \<U> \<approx> {..<n} \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5919
         pairwise (separatedin (subtopology X (topspace X - T))) \<U> \<and> {} \<notin> \<U> \<and> \<Union>\<U> = topspace X - T"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5920
    using separation_by_closed_intermediates_eq_count [of X n T] Kuratowski_aux2 lcX hnX by auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5921
  with False show ?thesis
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5922
    using lepoll_connected_components_alt by fastforce
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5923
qed auto
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5924
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5925
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5926
theorem Kuratowski_component_number_invariance:
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5927
  assumes "compact_space X" "Hausdorff_space X" "locally_connected_space X" "hereditarily normal_space X"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5928
  shows "((\<forall>S T n.
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5929
              closedin X S \<and> closedin X T \<and>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5930
              (subtopology X S) homeomorphic_space (subtopology X T)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5931
              \<longrightarrow> (connected_components_of
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5932
                    (subtopology X (topspace X - S)) \<approx> {..<n::nat} \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5933
                   connected_components_of
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5934
                    (subtopology X (topspace X - T)) \<approx> {..<n::nat})) \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5935
           (\<forall>S T n.
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5936
              (subtopology X S) homeomorphic_space (subtopology X T)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5937
              \<longrightarrow> (connected_components_of
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5938
                    (subtopology X (topspace X - S)) \<approx> {..<n::nat} \<longleftrightarrow>
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5939
                   connected_components_of
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5940
                    (subtopology X (topspace X - T)) \<approx> {..<n::nat})))"
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5941
         (is "?lhs = ?rhs")
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5942
proof
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5943
  assume L: ?lhs 
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5944
  then show ?rhs
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5945
    apply (subst (asm) Kuratowski_aux1, use homeomorphic_space_sym in blast)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5946
    apply (subst Kuratowski_aux1, use homeomorphic_space_sym in blast)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5947
    apply (blast intro: Kuratowski_component_number_invariance_aux assms)
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5948
    done
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5949
qed blast
f44aec9a6894 Last of the HOL Light metric space imports, and some supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  5950
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  5951
end