src/HOL/Analysis/Urysohn.thy
author paulson <lp15@cam.ac.uk>
Tue, 04 Jul 2023 12:53:01 +0100
changeset 78250 400aecdfd71f
parent 78248 740b23f1138a
child 78256 71e1aa0d9421
permissions -rw-r--r--
Another tranche of HOL Light material on metric and topological spaces
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
lemma Urysohn_lemma:
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
lemma Tietze_extension_closed_real_interval:
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"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
      by (simp add: abs_le_iff continuous_map_def minus_le_iff)
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
lemma Tietze_extension_realinterval:
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
subsection \<open>random metric space stuff\<close>
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
lemma metrizable_imp_k_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
  assumes "metrizable_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
  shows "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
  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
   676
    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
   677
  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
   678
    by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
    unfolding k_space Xeq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
  proof clarsimp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
    fix S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
    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
   684
    have "l \<in> S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
      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
   686
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
      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
   688
      interpret Submetric M d K
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
      proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
        show "K \<subseteq> M"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
          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
   692
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
      have "compactin mtopology K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
        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
   695
        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
   696
      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
   697
        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
   698
      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
   699
        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
   700
      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
   701
        using l 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
        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
   703
        by (force simp: K_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
      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
   705
        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
   706
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
        by simp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
    then show "closedin mtopology S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
      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
   711
      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
   712
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
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
   716
  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
   717
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
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
   719
  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
   720
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
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
   723
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
lemma hereditarily_B:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
  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
   726
      \<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
   727
  shows "hereditarily normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
  unfolding hereditarily_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
  fix W
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
  assume "W \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
  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
   733
    unfolding normal_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
    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
   737
      and "disjnt S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
    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
   739
      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
   740
    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
   741
      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
   742
    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
   743
      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
   744
      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
   745
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
lemma hereditarily_C: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
  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
   750
  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
   751
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
  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
   753
  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
   754
    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
   755
  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
   756
    unfolding ST_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
    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
   758
  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
   759
    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
   760
  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
   761
                         (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
   762
    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
   763
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
    using sub subX
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
    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
   766
    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
   767
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
lemma hereditarily_normal_space: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
  "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
   771
  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
   772
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
lemma hereditarily_normal_separation:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
  "hereditarily normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
        (\<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
   776
             \<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
   777
  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
   778
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
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
   781
   "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
   782
  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
   783
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
lemma metrizable_space_separation:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
   "\<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
   786
    \<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
   787
  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
   788
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
lemma hereditarily_normal_separation_pairwise:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
   "hereditarily normal_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
    (\<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
   792
        \<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
   793
                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
   794
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
  assume L: ?lhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
  show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
    fix \<U>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
    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
   801
      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
   802
    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
   803
                    (\<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
   804
      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
   805
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
      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
   807
        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
   808
      with L show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
        unfolding hereditarily_normal_separation
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
        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
   811
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
    then obtain \<F> \<G> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
        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
   814
        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
   815
        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
   816
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
    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
   818
    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
   819
    proof (intro exI conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
      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
   821
        unfolding \<H>_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
        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
   823
      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
   824
        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
   825
    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
   826
      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
   827
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
  assume R: ?rhs 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
    unfolding hereditarily_normal_separation
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
    fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
    assume "separatedin X S T"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
    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
   837
    proof (cases "T=S")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
        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
   841
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
      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
   844
        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
   845
      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
   846
        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
   847
        ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
        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
   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
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
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
   854
   "\<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
   855
  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
   856
  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
   857
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
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
   859
  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
   860
  shows  "hereditarily normal_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
  unfolding hereditarily
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
  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
   863
  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
   864
    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
   865
  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
   866
    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
   867
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
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
   871
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
definition completely_regular_space where
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
 "completely_regular_space X \<equiv>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
    \<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
   875
          \<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
   876
                  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
   877
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
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
   879
  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
   880
  shows "completely_regular_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
  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
   883
    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
   884
    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
   885
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
    unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
  proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
    fix S x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
    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
   890
    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
   891
      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
   892
    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
   893
      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
   894
    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
   895
      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
   896
    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
   897
      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
   898
    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
   899
      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
   900
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
lemma homeomorphic_completely_regular_space:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  assumes "X homeomorphic_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
  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
   906
  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
   907
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
lemma completely_regular_space_alt:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
   "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
     (\<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
   911
           \<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
   912
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
  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
   914
    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
   915
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
    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
   918
      using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
      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
   920
  qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
  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
   922
    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
   923
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
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
   926
lemma completely_regular_space_alt':
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
   "completely_regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
     (\<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
   929
           \<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
   930
  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
   931
  by (meson openin_subset subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
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
   934
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
  assumes "a \<noteq> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
  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
   937
         (\<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
   938
               \<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
   939
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
  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
   941
    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
   942
        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
   943
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
    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
   946
      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
   947
  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
   948
  moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
  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
   950
    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
   951
        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
   952
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
    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
   955
      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
   956
  qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
    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
   959
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
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
   962
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
   963
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
  assumes "a \<noteq> b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
  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
   966
         (\<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
   967
               \<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
   968
  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
   969
  by (meson openin_subset subsetD)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
lemma completely_regular_space_gen:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
  fixes a b::real
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
  assumes "a < b"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
  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
   975
         (\<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
   976
               \<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
   977
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
  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
   979
    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
   980
      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
   981
    for S x f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
    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
   984
      using that assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
      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
   986
  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
   987
  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
   988
    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
   989
    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
   990
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
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
   993
  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
   994
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
  unfolding completely_regular_space_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
  fix x S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
  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
   999
  assume "x \<in> topspace X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
  then have "closedin X {x}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
    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
  1002
  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
  1003
    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
  1004
  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
  1005
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
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
  1009
  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
  1010
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
  unfolding completely_regular_space_alt
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
  fix x S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
  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
  1015
  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
  1016
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
    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
  1018
  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
  1019
    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
  1020
    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
  1021
  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
  1022
    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
  1023
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
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
  1026
   "\<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
  1027
  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
  1028
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
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
  1030
   "\<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
  1031
  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
  1032
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
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
  1034
   "completely_regular_space mtopology"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
  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
  1036
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
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
  1038
   "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
  1039
  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
  1040
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
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
  1042
   "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
  1043
  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
  1044
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
lemma completely_regular_space_subtopology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
  assumes "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
  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
  1048
  unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
  fix A x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
  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
  1052
  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
  1053
    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
  1054
  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
  1055
    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
  1056
    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
  1057
    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
  1058
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
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
  1061
   " \<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
  1062
  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
  1063
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
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
  1065
  assumes "completely_regular_space X" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
  shows "regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
  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
  1069
    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
  1070
      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
  1071
    for C a f
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
    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
  1074
      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
  1075
      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
  1076
    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
  1077
      using a f0 by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
    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
  1079
      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
  1080
  qed (auto simp: disjnt_iff)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
    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
  1084
    by (meson "*")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
lemma 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
  1089
  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
  1090
  shows "completely_regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
  unfolding completely_regular_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
  fix S x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
  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
  1095
  have "regular_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
    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
  1097
  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
  1098
    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
  1099
  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
  1100
    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
  1101
  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
  1102
    by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
  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
  1104
    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
  1105
  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
  1106
    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
  1107
  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
  1108
    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
  1109
  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
  1110
    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
  1111
  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
  1112
    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
  1113
  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
  1114
    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
  1115
  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
  1116
    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
  1117
    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
  1118
  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
  1119
    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
  1120
    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
  1121
  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
  1122
  proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
    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
  1124
      unfolding continuous_map_closedin
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
    proof (intro strip conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
      fix C
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
      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
  1128
      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
  1129
        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
  1130
      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
  1131
        unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
      proof (intro closedin_Un)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
        have "closedin euclidean C"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
          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
  1135
        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
  1136
          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
  1137
        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
  1138
          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
  1139
      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
  1140
    qed (use gim in force)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
    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
  1142
      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
  1143
    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
  1144
      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
  1145
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
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
  1149
   "locally_compact_space X
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
        \<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
  1151
  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
  1152
  by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
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
  1155
   "completely_regular_space (prod_topology X Y) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
      topspace (prod_topology X Y) = {} \<or>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
      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
  1158
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
    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
  1161
       (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
  1162
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
  proof (cases "topspace(prod_topology X Y) = {}")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
    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
  1168
      using R by blast+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
      unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
    proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
      fix W x y
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
      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
  1174
      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
  1175
        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
  1176
      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
  1177
        using openin_subset by fastforce+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
      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
  1179
        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
  1180
        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
  1181
        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
  1182
      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
  1183
        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
  1184
        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
  1185
        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
  1186
      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
  1187
      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
  1188
      proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
        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
  1190
          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
  1191
        moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
        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
  1193
          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
  1194
        ultimately
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
        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
  1196
          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
  1197
          by (intro continuous_intros) auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
        show "h (x, y) = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
          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
  1200
        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
  1201
          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
  1202
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
  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
  1205
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
lemma completely_regular_space_product_topology:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
   "completely_regular_space (product_topology X I) \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
    (\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. completely_regular_space (X i))" 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
    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
  1215
       (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
  1216
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
  proof (cases "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {}")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
      unfolding completely_regular_space_alt'
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
    proof clarify
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
      fix W x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
      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
  1226
      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
  1227
             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
  1228
             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
  1229
        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
  1230
      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
  1231
              \<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
  1232
                       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
  1233
        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
  1234
        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
  1235
      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
  1236
              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
  1237
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
      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
  1239
                                             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
  1240
        by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
      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
  1242
      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
  1243
                     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
  1244
      proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
        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
  1246
          using f that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
          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
  1248
        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
  1249
          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
  1250
        show "h x = 0"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
          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
  1252
        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
  1253
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
          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
  1255
            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
  1256
            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
  1257
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
            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
  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      
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
  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
  1263
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1266
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
  1267
  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
  1268
  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
  1269
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
  1270
  fix C a
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1271
  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
  1272
  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
  1273
    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
  1274
  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
  1275
  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
  1276
    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
  1277
      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
  1278
      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
  1279
    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
  1280
      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
  1281
  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
  1282
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1283
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1284
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
  1285
  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
  1286
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
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
  1288
   "t1_space mtopology"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
  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
  1290
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
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
  1293
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
definition kification_open 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
  where "kification_open \<equiv> 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
          \<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
  1297
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
definition kification 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
  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
  1300
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
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
  1302
  unfolding istopology_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
  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
  1305
    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
  1306
    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
  1307
    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
  1308
  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
  1309
    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
  1310
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
lemma openin_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
   "openin (kification X) U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
        U \<subseteq> topspace X \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
        (\<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
  1316
  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
  1317
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
lemma openin_kification_finer:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
   "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
  1320
  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
  1321
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
lemma topspace_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
   "topspace(kification X) = topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
  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
  1325
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
lemma closedin_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
   "closedin (kification X) U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
      U \<subseteq> topspace X \<and>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
      (\<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
  1330
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
  1331
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
    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
  1334
qed (simp add: closedin_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
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
  1337
  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
  1338
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
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
  1340
  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
  1341
  by (metis openin_closedin_eq)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
lemma compactin_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
   "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
  1345
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
    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
  1348
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
    unfolding compactin_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
    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
  1354
      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
  1355
    fix \<U>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
    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
  1357
    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
  1358
      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
  1359
    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
  1360
      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
  1361
    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
  1362
      unfolding compact_space
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
      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
  1364
    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
  1365
      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
  1366
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
lemma compact_space_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
   "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
  1371
  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
  1372
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
lemma kification_kification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
   "kification(kification X) = kification X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
  unfolding openin_inject [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
  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
  1379
  proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
    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
  1381
      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
  1382
  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
  1383
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
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
  1386
  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
  1387
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
lemma continuous_map_into_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
  assumes "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
  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
  1391
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
    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
  1394
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
  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
  1397
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
    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
  1399
      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
  1400
    proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
      have "compactin Y (f ` K)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
        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
  1403
      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
  1404
        by (meson V openin_kification)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
      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
  1406
        by (meson openin_subtopology)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
        unfolding openin_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
      proof (intro conjI exI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
        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
  1411
          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
  1412
      qed (use U in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
      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
  1416
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
  with R show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
    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
  1419
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
lemma continuous_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
   "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
  1423
  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
  1424
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
lemma continuous_map_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
   "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
  1427
  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
  1428
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
lemma subtopology_kification_compact:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
  assumes "compactin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
  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
  1432
  unfolding openin_inject [symmetric]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
  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
  1436
    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
  1437
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
lemma subtopology_kification_finer:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
  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
  1442
  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
  1443
  using assms 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
  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
  1445
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
lemma proper_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
  assumes "k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
  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
  1449
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
  assume ?lhs then show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
    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
  1452
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
  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
  1455
    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
  1456
  with R show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
    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
  1458
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
lemma perfect_map_from_kification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
   "\<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
  1462
  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
  1463
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
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
  1465
  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
  1466
  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
  1467
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
  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
  1469
    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
  1470
  assume "k_space Y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
  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
  1472
    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
  1473
    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
  1474
  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
  1475
    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
  1476
  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
  1477
    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
  1478
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
  1479
  then show "k_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
    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
  1481
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
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
  1485
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
lemma one_point_compactification_dense:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
   "\<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
  1488
  unfolding closure_of_complement
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
  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
  1490
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
lemma one_point_compactification_interior:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
   "\<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
  1493
  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
  1494
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
lemma kc_space_one_point_compactification_gen:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
  shows "kc_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
         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
  1499
         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
  1500
 (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
  assume L: ?lhs show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
    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
  1505
      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
  1506
    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
  1507
      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
  1508
    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
  1509
      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
  1510
    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
  1511
      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
  1512
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
    unfolding kc_space_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
    fix S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
    assume "compactin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
    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
  1521
      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
  1522
    show "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
    proof (cases "a \<in> S")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
      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
  1526
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
      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
  1528
      proof (rule openin_trans_full)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
        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
  1530
        proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
          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
  1532
            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
  1533
          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
  1534
            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
  1535
          proof (intro closedin_subset_topspace)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
            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
  1537
              using that
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
              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
  1539
          qed (use that in auto)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
          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
  1541
            using R by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
          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
  1543
            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
  1544
          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
  1545
            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
  1546
            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
  1547
        qed 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
        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
  1549
          by (simp add: R)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
      ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
        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
  1553
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
        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
  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
qed
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
  
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
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
  1563
  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
  1564
| 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
  1565
78128
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1566
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
  1567
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
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
  1569
   (\<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
  1570
  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
  1571
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
lemma Alexandroff_open_Un_aux:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
  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
  1574
  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
  1575
  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
  1576
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
  1577
  case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
    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
  1580
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
  have "U \<subseteq> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
    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
  1584
  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
  1585
    by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
  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
  1587
    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
  1588
  moreover
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
  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
  1590
    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
  1591
  ultimately show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
    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
  1593
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
lemma Alexandroff_open_Un:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
  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
  1597
  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
  1598
  using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
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
  1600
  case (base U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
    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
  1603
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
    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
  1607
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
lemma Alexandroff_open_Int_aux:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
  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
  1611
  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
  1612
  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
  1613
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
  1614
  case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
    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
  1617
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
  case (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
  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
  1620
    by force
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
  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
  1622
    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
  1623
  then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
    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
  1625
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
lemma istopology_Alexandroff_open: "istopology (Alexandroff_open X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
  unfolding istopology_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
  fix S T
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
  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
  1632
  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
  1633
  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
  1634
    case (base U)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
    then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
      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
  1637
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
    case EC: (ext C)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
    show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
      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
  1641
    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
  1642
      case (base V)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
      then show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
        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
  1645
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
      case (ext D)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
      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
  1648
              = 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
  1649
        by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
      show ?case
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
        unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
        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
  1653
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
  fix \<K>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
  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
  1658
  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
  1659
  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
  1660
    case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
    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
  1662
      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
  1663
    then obtain U where U: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
      "\<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
  1665
                    \<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
  1666
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
    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
  1668
    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
  1669
    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
  1670
    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
  1671
      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
  1672
    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
  1673
      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
  1674
    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
  1675
    proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
      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
  1677
        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
  1678
      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
  1679
        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
  1680
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
    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
  1682
      using U2 True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
      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
  1684
    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
  1685
      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
  1686
    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
  1687
      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
  1688
    show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
      unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
    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
  1691
      show "openin X A"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
        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
  1693
      show "closedin X B"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
        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
  1695
      show "compactin X B"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
        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
  1697
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
    case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
    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
  1701
      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
  1702
    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
  1703
      by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
    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
  1705
      using image_iff by fastforce
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
    show ?thesis
78128
3d2db8057b9f Hiding the constructor names, particularly to avoid conflicts involving "ext"
paulson <lp15@cam.ac.uk>
parents: 78127
diff changeset
  1707
      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
  1708
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
definition Alexandroff_compactification where
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
  "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
  1714
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
lemma openin_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
   "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
  1717
        (\<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
  1718
        (\<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
  1719
  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
  1720
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
lemma topspace_Alexandroff_compactification [simp]:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
   "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
  1724
   (is "?lhs = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
  show "?lhs \<subseteq> ?rhs"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1727
    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
  1728
  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
  1729
    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
  1730
  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
  1731
    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
  1732
    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
  1733
  ultimately show "?rhs \<subseteq> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
    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
  1735
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
lemma closedin_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
   "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
  1739
      (\<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
  1740
      (\<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
  1741
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
    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
  1745
    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
  1746
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
    using openin_subset 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
    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
  1749
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
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
  1752
   "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
  1753
  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
  1754
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
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
  1756
   "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
  1757
  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
  1758
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
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
  1760
  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
  1761
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
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
  1763
  unfolding continuous_map_def 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
  fix U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
  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
  1767
  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
  1768
    | 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
  1769
    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
  1770
  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
  1771
  proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
    case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
      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
  1775
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
    case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
      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
  1779
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
qed auto
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
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
  1784
  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
  1785
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
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
  1787
   "compact_space(Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
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
  1789
  fix \<U> U
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
  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
  1791
      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
  1792
      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
  1793
  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
  1794
    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
  1795
  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
  1796
  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
  1797
          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
  1798
    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
  1799
  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
  1800
    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
  1801
  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
  1802
    | 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
  1803
    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
  1804
  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
  1805
  proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
    case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
    then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
      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
  1809
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
    case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
    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
  1812
      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
  1813
    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
  1814
      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
  1815
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
lemma topspace_Alexandroff_compactification_delete:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
   "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
  1820
  by simp
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
lemma Alexandroff_compactification_dense:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
  assumes "\<not> compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
  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
  1825
         topspace(Alexandroff_compactification X)"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
  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
  1827
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
  1828
  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
  1829
    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
  1830
qed auto
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
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
  1834
  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
  1835
  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
  1836
   (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
    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
  1840
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
    using assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
    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
  1843
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
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
  1846
   "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
  1847
  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
  1848
  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
  1849
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
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
  1851
  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
  1852
    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
  1853
  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
  1854
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
    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
  1857
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
    unfolding t1_space_closedin_singleton
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
    fix x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
    assume "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
    show "closedin X {x}"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
    proof (cases "x=a")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
        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
  1868
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1871
        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
  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
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
lemma closedin_Alexandroff_I: 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
  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
  1878
          "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
  1879
  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
  1880
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
  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
  1882
    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
  1883
  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
  1884
    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
  1885
  moreover have "closedin X S"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
    using assms S
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
    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
  1888
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
    by (simp add: S)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
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
  1894
  "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
  1895
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
  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
  1897
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
    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
  1900
    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
  1901
    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
  1902
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
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
  1906
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
    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
  1908
    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
  1909
                \<Longrightarrow> closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
  shows "kc_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
         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
  1912
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
  have "closedin X K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
    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
  1915
    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
  1916
    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
  1917
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
    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
  1919
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
lemma kc_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
  "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
  1923
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
  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
  1925
      (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
  1926
    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
  1927
  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
  1928
    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
  1929
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
lemma regular_space_one_point_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
  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
  1935
    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
  1936
  shows "regular_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
           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
  1938
    (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1940
  show "?lhs \<Longrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
    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
  1942
  assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
  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
  1944
  show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
    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
  1946
  proof (intro strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
    fix W x
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
    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
  1949
    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
  1950
    proof (cases "x=a")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
      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
  1953
        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
  1954
        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
  1955
      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
  1956
        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
  1957
      show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
      proof (intro exI conjI)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
        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
  1960
          using "\<section>" K by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
        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
  1962
          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
  1963
        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
  1964
        proof (rule)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
          show "x \<in> topspace X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
            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
  1967
          show "x \<notin> K"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
            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
  1969
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
        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
  1971
          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
  1972
        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
  1973
          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
  1974
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
      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
  1978
        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
  1979
      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
  1980
        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
  1981
      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
  1982
               "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
  1983
        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
  1984
        by (metis (no_types, lifting))
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
        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
  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
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
lemma regular_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
  "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
  1994
  (is "regular_space ?Y = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
  have "regular_space ?Y \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
        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
  1998
    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
  1999
  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
  2000
    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
  2001
        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
  2002
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
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
  2007
  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
  2008
    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
  2009
  shows "Hausdorff_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
           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
  2011
    (is "?lhs \<longleftrightarrow> ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
  show ?rhs if ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
  proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
    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
  2016
      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
  2017
      by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
    with that show ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
      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
  2020
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
  show "?rhs \<Longrightarrow> ?lhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
    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
  2024
        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
  2025
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
lemma Hausdorff_space_Alexandroff_compactification:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
   "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
  2029
  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
  2030
      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
  2031
      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
  2032
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
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
  2034
   "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
  2035
        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
  2036
  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
  2037
      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
  2038
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
lemma Hausdorff_space_one_point_compactification_asymmetric_prod:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
  assumes "compact_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
  shows "Hausdorff_space X \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
         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
  2043
         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
  2044
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
  2045
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
  show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
  proof 
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
    show ?rhs if ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
    proof
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
      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
  2051
        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
  2052
      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
  2053
        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
  2054
            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
  2055
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
  next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
    assume R: ?rhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
    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
  2059
    show ?lhs
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
    proof (cases "topspace X = {a}")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
        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
  2064
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
      then have "kc_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2067
        using kc_space_retraction_map_image [of "prod_topology X (subtopology X (topspace X - {a}))" X fst]
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2068
        by (metis Diff_subset R True insert_Diff retraction_map_fst topspace_subtopology_subset)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
      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
  2070
        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
  2071
      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
  2072
        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
  2073
          by force+
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
        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
  2075
          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
  2076
        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
  2077
          unfolding eq
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
        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
  2079
          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
  2080
            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
  2081
          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
  2082
            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
  2083
          ultimately
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
          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
  2085
            unfolding compactin_subtopology
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
            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
  2087
          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
  2088
            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
  2089
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
        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
  2091
          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
  2092
      qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
      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
  2094
        by (auto simp: D_def)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
      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
  2096
        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
  2097
      have "x=y"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
        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
  2099
          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
  2100
      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
  2101
        case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
        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
  2103
          by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
        then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
        proof cases
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
          case 1
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
          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
  2108
            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
  2109
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
            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
  2111
              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
  2112
            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
  2113
            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
  2114
              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
  2115
                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
  2116
              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
  2117
                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
  2118
              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
  2119
                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
  2120
                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
  2121
              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
  2122
                by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
            qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
            ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
              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
  2126
          qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
          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
  2128
            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
  2129
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
            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
  2131
        next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
          case 2
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
          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
  2134
            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
  2135
          proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
            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
  2137
              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
  2138
            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
  2139
              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
  2140
              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
  2141
            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
  2142
              by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
            then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
              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
  2145
          qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
          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
  2147
            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
  2148
          then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
            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
  2150
        qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
      qed auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
        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
  2154
          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
  2155
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
  case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
    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
  2161
qed
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
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2164
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
  2165
   "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
  2166
        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
  2167
        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
  2168
    (is "Hausdorff_space ?Y = ?rhs")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
  have *: "subtopology (Alexandroff_compactification X)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
     (topspace (Alexandroff_compactification X) -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
      {None}) homeomorphic_space X"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
    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
  2174
  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
  2175
      (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
  2176
       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
  2177
    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
  2178
  also have "... \<longleftrightarrow> ?rhs"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
    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
  2180
          homeomorphic_space_refl * by blast
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
  finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2182
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2183
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
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
  2185
  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
  2186
  shows "openin X U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
         (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
  2188
                   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
  2189
proof (cases "a \<in> U")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
  case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
    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
  2193
next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
  case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
  then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
  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
  2197
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
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
  2200
  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
  2201
  shows "openin X U \<longleftrightarrow>
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
         (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
  2203
                     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
  2204
                     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
  2205
                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
  2206
  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
  2207
  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
  2208
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
lemma Alexandroff_compactification_unique:
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
  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
  2211
  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
  2212
        (is "?Y homeomorphic_space X")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
proof -
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
  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
  2215
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
  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
  2217
               "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
  2218
    by auto
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
  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
  2220
    unfolding quotient_map_def
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
  proof (intro conjI strip)
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
    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
  2223
      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
  2224
    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
  2225
      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
  2226
    proof (cases "None \<in> U")
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
      case True
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
      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
  2229
        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
  2230
      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
  2231
        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
  2232
      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
  2233
        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
  2234
      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
  2235
        by metis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
      also have "\<dots> \<longleftrightarrow> ?R"
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
        using Tsub assms
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
        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
  2239
        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
  2240
      finally show ?thesis .
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
    next
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
      case False
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
      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
  2244
        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
  2245
      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
  2246
        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
  2247
      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
  2248
        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
  2249
      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
  2250
        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
  2251
      then show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
        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
  2253
    qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
  qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
  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
  2256
    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
  2257
  ultimately show ?thesis
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
    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
  2259
qed
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2260
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2261
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
  2262
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2263
lemma continuous_map_on_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2264
  assumes Y: "regular_space Y" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2265
    and T: "T \<subseteq> X closure_of S" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2266
    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
  2267
  shows "continuous_map (subtopology X T) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2268
proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2269
  fix a
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2270
  assume "a \<in> topspace X" and "a \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2271
  have "f ` T \<subseteq> topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2272
    by (metis f image_subsetI limitin_topspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2273
  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
  2274
    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
  2275
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2276
    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
  2277
      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
  2278
    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
  2279
      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
  2280
    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
  2281
      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
  2282
    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
  2283
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2284
      have "z \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2285
        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
  2286
      then have "f z \<in> topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2287
        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
  2288
      { assume "f z \<in> topspace Y" "f z \<notin> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2289
        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
  2290
          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
  2291
        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
  2292
                 "\<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
  2293
          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
  2294
        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
  2295
          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
  2296
        have False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2297
          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
  2298
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2299
      }
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2300
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2301
        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
  2302
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2303
    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
  2304
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2305
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2306
      using eventually_atin_within by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2307
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2308
  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
  2309
    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
  2310
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2313
lemma continuous_map_on_intermediate_closure_of_eq:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2314
  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
  2315
  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
  2316
        (is "?lhs \<longleftrightarrow> ?rhs")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2317
proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2318
  assume L: ?lhs
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2319
  show ?rhs
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2320
  proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2321
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2322
    assume "x \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2323
    with L Tsub closure_of_subset_topspace 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2324
    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
  2325
      by (fastforce simp: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2326
    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
  2327
      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
  2328
      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
  2329
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2330
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2331
  show "?rhs \<Longrightarrow> ?lhs" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2332
    using assms continuous_map_on_intermediate_closure_of by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2333
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2334
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2335
lemma continuous_map_extension_pointwise_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2336
  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
  2337
    and f: "continuous_map (subtopology X S) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2338
    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
  2339
  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
  2340
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2341
  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
  2342
    by (metis Diff_iff lim)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2343
  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
  2344
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2345
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2346
    have T: "T \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2347
      using \<section> closure_of_subset_topspace by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2348
    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
  2349
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2350
      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
  2351
        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
  2352
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2353
        by (simp add: eventually_atin_within limitin_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2354
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2355
    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
  2356
      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
  2357
    ultimately show "continuous_map (subtopology X T) Y ?h"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2358
      unfolding continuous_map_on_intermediate_closure_of_eq [OF \<section>] 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2359
      by (auto simp: \<section> atin_subtopology_within)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2360
  qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2361
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2364
lemma continuous_map_extension_pointwise:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2365
  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
  2366
    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
  2367
                     (\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2368
  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
  2369
proof (rule continuous_map_extension_pointwise_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2370
  show "continuous_map (subtopology X S) Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2371
  proof (clarsimp simp add: continuous_map_atin)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2372
    fix t
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2373
    assume "t \<in> topspace X" and "t \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2374
    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
  2375
      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
  2376
    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
  2377
      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
  2378
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2379
  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
  2380
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2381
    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
  2382
      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
  2383
    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
  2384
      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
  2385
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2386
      unfolding limitin_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2387
      by (smt (verit) eventually_atin_within gf subsetD subset_insertI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2388
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2389
qed (use assms in auto)
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2392
subsection\<open>Extending Cauchy continuous functions to the closure\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2393
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2394
lemma Cauchy_continuous_map_extends_to_continuous_closure_of_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2395
  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
  2396
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2397
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2398
  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
  2399
                        (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
  2400
proof (rule continuous_map_extension_pointwise_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2401
  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
  2402
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2403
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2404
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2405
  show "regular_space (mtopology_of m2)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2406
    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
  2407
  show "S \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2408
    by (simp add: assms(3) closure_of_subset)    
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2409
  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
  2410
    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
  2411
  fix a
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2412
  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
  2413
  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
  2414
    and lim\<sigma>: "limitin L.M1.mtopology \<sigma> a sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2415
    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
  2416
  then have "L.M1.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2417
    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
  2418
  then have "L.M2.MCauchy (f \<circ> \<sigma>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2419
    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
  2420
  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
  2421
    by (meson L.M2.mcomplete_def m2 mcomplete_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2422
  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
  2423
    unfolding L.limit_atin_sequentially_within imp_conjL
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2424
  proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2425
    show "l \<in> mspace m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2426
      using L.M2.limitin_mspace l by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2427
    fix \<rho>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2428
    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
  2429
    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
  2430
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2431
    have "a \<in> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2432
      using L.M1.limitin_mspace lim\<rho> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2433
    have "S.MCauchy \<sigma>" "S.MCauchy \<rho>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2434
      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
  2435
    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
  2436
      using f by (auto simp: Cauchy_continuous_map_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2437
    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
  2438
      by (auto simp: L.M2.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2439
    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
  2440
    proof (rule Lim_null_comparison)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2441
      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
  2442
        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
  2443
      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
  2444
        by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2445
      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
  2446
      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
  2447
      proof (rule Lim_null_comparison)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2448
        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
  2449
          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
  2450
        have "(\<lambda>n. mdist m1 (\<sigma> n) a) \<longlonglongrightarrow> 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2451
          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
  2452
        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
  2453
          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
  2454
        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
  2455
          by (simp add: tendsto_add_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2456
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2457
      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
  2458
        by (simp add: S.MCauchy_interleaving_gen \<psi>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2459
      then have "L.M2.MCauchy (f \<circ> \<psi>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2460
        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
  2461
      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
  2462
        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
  2463
        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
  2464
      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
  2465
        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
  2466
      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
  2467
        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
  2468
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2469
    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
  2470
      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
  2471
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2472
  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
  2473
    by (force simp: mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2474
qed auto
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2477
lemma Cauchy_continuous_map_extends_to_continuous_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2478
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2479
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2480
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2481
  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
  2482
                        (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
  2483
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2484
  obtain g where cmg: 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2485
    "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
  2486
                        (mtopology_of m2) g" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2487
    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
  2488
    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
  2489
    by (metis inf_commute inf_le2 submetric_restrict)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2490
  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
  2491
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2492
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2493
    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
  2494
                         (mtopology_of m2) h"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2495
      unfolding h_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2496
    proof (rule continuous_map_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2497
      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
  2498
        by (metis closure_of_restrict cmg topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2499
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2500
  qed (auto simp: gf h_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2501
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2504
lemma Cauchy_continuous_map_extends_to_continuous_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2505
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2506
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2507
    and T: "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2508
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2509
  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
  2510
         "(\<forall>x \<in> S. g x = f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2511
  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
  2512
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2513
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
  2514
lemma all_in_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2515
  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
  2516
  shows "\<forall>x \<in> X closure_of S. P x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2517
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2518
  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
  2519
    using P by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2520
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2521
  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
  2522
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2523
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2524
lemma Lipschitz_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2525
  assumes lcf: "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2526
    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
  2527
    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
  2528
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2529
  shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2530
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2531
  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
  2532
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2533
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2534
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2535
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2536
    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
  2537
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2538
    unfolding Lipschitz_continuous_map_pos
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2539
  proof
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  2540
    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
  2541
      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
  2542
          mtopology_of_submetric image_subset_iff_funcset)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2543
    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
  2544
    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
  2545
      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
  2546
    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
  2547
        for p q and A::"('a*'a)set"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2548
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2549
    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
  2550
      unfolding eq
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2551
    proof (rule closedin_continuous_map_preimage)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2552
      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
  2553
        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
  2554
      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
  2555
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2556
      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
  2557
        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
  2558
          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
  2559
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2560
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2561
    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
  2562
      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
  2563
      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
  2564
          mtopology_of_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2565
    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
  2566
             \<forall>y\<in>mspace (submetric m1 T).
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2567
                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
  2568
      using \<open>0 < B\<close> by auto
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
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2573
lemma Lipschitz_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2574
  assumes "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2575
    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
  2576
    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
  2577
  shows "Lipschitz_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2578
  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
  2579
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2580
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2581
lemma Lipschitz_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2582
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2583
    and f: "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2584
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2585
  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
  2586
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2587
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2588
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2589
    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
  2590
                        (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
  2591
    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
  2592
  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
  2593
  proof (rule Lipschitz_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2594
    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
  2595
      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
  2596
    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
  2597
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2598
    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
  2599
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2600
    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
  2601
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2602
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2603
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2604
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2605
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2608
lemma Lipschitz_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2609
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2610
    and "Lipschitz_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2611
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2612
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2613
  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
  2614
  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
  2615
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2616
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
  2617
lemma uniformly_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2618
  assumes ucf: "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2619
    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
  2620
    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
  2621
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2622
  shows "uniformly_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2623
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2624
  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
  2625
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2626
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2627
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2628
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2629
    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
  2630
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2631
    unfolding uniformly_continuous_map_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2632
    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
  2633
    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
  2634
      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
  2635
          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
  2636
    fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2637
    assume "\<epsilon> > 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2638
    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
  2639
      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
  2640
      apply (simp add: Ball_def del: divide_const_simps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2641
      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
  2642
    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
  2643
    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
  2644
    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
  2645
            = {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
  2646
      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
  2647
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2648
    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
  2649
      unfolding eq
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2650
    proof (intro closedin_Un closedin_continuous_map_preimage)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2651
      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
  2652
        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
  2653
      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
  2654
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2655
        by (intro continuous_intros; simp add: mtopology_of_def *)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2656
      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
  2657
        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
  2658
      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
  2659
        unfolding case_prod_unfold
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2660
        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
  2661
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2662
    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
  2663
      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
  2664
      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
  2665
          mtopology_of_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2666
    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
  2667
      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
  2668
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2669
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2670
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2671
lemma uniformly_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2672
  assumes "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2673
    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
  2674
    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
  2675
  shows "uniformly_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2676
  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
  2677
      uniformly_continuous_map_on_intermediate_closure_aux)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2678
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2679
lemma uniformly_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2680
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2681
    and f: "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2682
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2683
  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
  2684
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2685
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2686
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2687
    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
  2688
                        (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
  2689
    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
  2690
  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
  2691
  proof (rule uniformly_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2692
    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
  2693
      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
  2694
    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
  2695
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2696
    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
  2697
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2698
    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
  2699
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2700
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2701
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2702
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2703
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2706
lemma uniformly_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2707
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2708
    and "uniformly_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2709
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2710
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2711
  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
  2712
  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
  2713
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2714
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2715
lemma Cauchy_continuous_map_on_intermediate_closure_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2716
  assumes ucf: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2717
    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
  2718
    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
  2719
    and "S \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2720
  shows "Cauchy_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2721
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2722
  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
  2723
    by (simp add: Metric_space12_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2724
  interpret S: Metric_space "S \<inter> mspace m1" "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2725
    by (simp add: L.M1.subspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2726
  interpret T: Metric_space T "mdist m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2727
    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
  2728
  have "T \<subseteq> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2729
    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
  2730
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2731
  proof (clarsimp simp: Cauchy_continuous_map_def Int_absorb2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2732
    fix \<sigma>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2733
    assume \<sigma>: "T.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2734
    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
  2735
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2736
      have "\<sigma> n \<in> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2737
        using \<sigma> by (force simp: T.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2738
      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
  2739
        by (metis cmf mtopology_of_def mtopology_of_submetric)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2740
      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
  2741
        using \<open>T \<subseteq> mspace m1\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2742
        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
  2743
        by (metis inverse_Suc of_nat_Suc)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2744
      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
  2745
        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
  2746
        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
  2747
        by (smt (verit, del_insts) inverse_Suc )
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2748
      with \<delta> \<open>S \<subseteq> T\<close> show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2749
        by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2750
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2751
    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
  2752
                    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
  2753
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2754
    have "S.MCauchy \<rho>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2755
      unfolding S.MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2756
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2757
      show "range \<rho> \<subseteq> S \<inter> mspace m1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2758
        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
  2759
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2760
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2761
      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
  2762
        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
  2763
      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
  2764
        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
  2765
      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
  2766
        by (meson eventually_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2767
      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
  2768
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2769
        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
  2770
          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
  2771
        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
  2772
          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
  2773
        also have "\<dots> < \<epsilon>/4 + \<epsilon>/2 + \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2774
          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
  2775
        also have "... \<le> \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2776
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2777
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2778
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2779
      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
  2780
        by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2781
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2782
    then have f\<rho>: "L.M2.MCauchy (f \<circ> \<rho>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2783
      using ucf by (simp add: Cauchy_continuous_map_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2784
    show "L.M2.MCauchy (f \<circ> \<sigma>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2785
      unfolding L.M2.MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2786
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2787
      show "range (f \<circ> \<sigma>) \<subseteq> mspace m2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2788
        using \<open>T \<subseteq> mspace m1\<close> \<sigma> cmf
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2789
        apply (auto simp: )
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2790
        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
  2791
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2792
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2793
      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
  2794
        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
  2795
      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
  2796
        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
  2797
      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
  2798
        by (meson eventually_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2799
      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
  2800
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2801
        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
  2802
          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
  2803
        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
  2804
          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
  2805
        also have "\<dots> < \<epsilon>/4 + \<epsilon>/2 + \<epsilon>/4"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2806
          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
  2807
        also have "... \<le> \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2808
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2809
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2810
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2811
      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
  2812
        by blast
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
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2816
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2817
lemma Cauchy_continuous_map_on_intermediate_closure:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2818
  assumes "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2819
    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
  2820
    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
  2821
  shows "Cauchy_continuous_map (submetric m1 T) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2822
  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
  2823
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2824
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2825
lemma Cauchy_continuous_map_extends_to_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2826
  assumes m2: "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2827
    and f: "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2828
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2829
  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
  2830
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2831
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2832
  obtain g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2833
    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
  2834
                        (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
  2835
    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
  2836
  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
  2837
  proof (rule Cauchy_continuous_map_on_intermediate_closure)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2838
    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
  2839
      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
  2840
    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
  2841
      using closure_of_subset_Int by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2842
    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
  2843
      by (metis closure_of_restrict subset_refl topspace_mtopology_of)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2844
    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
  2845
      by (simp add: g)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2846
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2847
  with g that show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2848
    by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2849
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2852
lemma Cauchy_continuous_map_extends_to_intermediate_closure_of:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2853
  assumes "mcomplete_of m2" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2854
    and "Cauchy_continuous_map (submetric m1 S) m2 f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2855
    and "T \<subseteq> mtopology_of m1 closure_of S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2856
  obtains g 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2857
  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
  2858
  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
  2859
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2860
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2861
subsection\<open>Metric space of bounded functions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2862
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2863
context Metric_space
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2864
begin
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2865
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2866
definition fspace :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) set" where 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2867
  "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
  2868
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2869
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
  2870
  "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
  2871
                    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
  2872
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2873
lemma fspace_empty [simp]: "fspace {} = {\<lambda>x. undefined}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2874
  by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2875
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2876
lemma fdist_empty [simp]: "fdist {} = (\<lambda>x y. 0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2877
  by (auto simp: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2878
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2879
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
  2880
  by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2881
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2882
lemma bdd_above_dist:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2883
  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
  2884
  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
  2885
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2886
  obtain a where "a \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2887
    using \<open>S \<noteq> {}\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2888
  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
  2889
    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
  2890
    using f g mbounded_pos by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2891
  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
  2892
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2893
    have "f u \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2894
      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
  2895
    have "g u \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2896
      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
  2897
    have "x \<in> M" "y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2898
      using B C that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2899
    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
  2900
      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
  2901
    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
  2902
      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
  2903
    also have "\<dots> \<le> B + d x y + C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2904
      using B C commute that by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2905
    finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2906
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2907
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2908
    by (meson bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2909
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2912
lemma Metric_space_funspace: "Metric_space (fspace S) (fdist S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2913
proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2914
  show *: "0 \<le> fdist S f g" for f g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2915
    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
  2916
  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
  2917
    by (auto simp: fdist_def commute)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2918
  show "(fdist S f g = 0) = (f = g)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2919
    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
  2920
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2921
    assume 0: "fdist S f g = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2922
    show "f = g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2923
    proof (cases "S={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2924
      case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2925
      with 0 that show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2926
        by (simp add: fdist_def fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2927
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2928
      case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2929
      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
  2930
        by (simp add: fdist_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2931
      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
  2932
          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
  2933
      with fg show "f=g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2934
        by (simp add: fspace_def extensionalityI image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2935
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2936
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2937
    show "f = g \<Longrightarrow> fdist S f g = 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2938
      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
  2939
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2940
  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
  2941
    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
  2942
  proof (clarsimp simp add: fdist_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2943
    assume "S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2944
    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
  2945
      by (meson fgh fspace_in_M that triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2946
    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
  2947
      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
  2948
    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
  2949
      by (auto simp: bdd_above_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2950
    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
  2951
      by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2952
    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
  2953
      by (auto simp: bdd_above_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2954
    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
  2955
      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
  2956
    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
  2957
      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
  2958
    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
  2959
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2960
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2961
78127
24b70433c2e8 New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2962
end
78202
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2965
definition funspace where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2966
  "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
  2967
                          Metric_space.fdist (mspace m) (mdist m) S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2968
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2969
lemma mspace_funspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2970
  "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
  2971
  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
  2972
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2973
lemma mdist_funspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2974
  "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
  2975
  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
  2976
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2977
lemma funspace_imp_welldefined:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2978
   "\<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
  2979
  by (simp add: Metric_space.fspace_def subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2980
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2981
lemma funspace_imp_extensional:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2982
   "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
  2983
  by (simp add: Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2984
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2985
lemma funspace_imp_bounded_image:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2986
   "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
  2987
  by (simp add: Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2988
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2989
lemma funspace_imp_bounded:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2990
   "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
  2991
  by (auto simp: Metric_space.fspace_def Metric_space.mbounded)
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2994
lemma (in Metric_space) funspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2995
  assumes "f \<in> fspace S" "g \<in> fspace S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2996
  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
  2997
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2998
  have "mbounded (f ` S \<union> g ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  2999
    using mbounded_Un assms by (force simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3000
  then show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3001
    by (metis UnCI imageI mbounded_alt that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3002
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3003
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3004
lemma funspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3005
  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
  3006
  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
  3007
  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
  3008
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3009
lemma (in Metric_space) funspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3010
  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
  3011
  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
  3012
    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
  3013
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3014
lemma funspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3015
  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
  3016
  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
  3017
  using assms by (simp add: Metric_space.funspace_mdist_le)
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3020
lemma (in Metric_space) mcomplete_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3021
  assumes "mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3022
  shows "mcomplete_of (funspace S Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3023
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3024
  interpret F: Metric_space "fspace S" "fdist S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3025
    by (simp add: Metric_space_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3026
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3027
  proof (cases "S={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3028
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3029
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3030
      by (simp add: mcomplete_of_def mcomplete_trivial_singleton)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3031
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3032
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3033
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3034
    proof (clarsimp simp: mcomplete_of_def Metric_space.mcomplete_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3035
      fix \<sigma>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3036
      assume \<sigma>: "F.MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3037
      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
  3038
        by (auto simp: F.MCauchy_def intro: fspace_in_M)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3039
      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
  3040
        using \<sigma> that by (auto simp: F.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3041
      have \<sigma>ext: "\<And>n. \<sigma> n \<in> extensional S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3042
        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
  3043
      have \<sigma>bd: "\<And>n. mbounded (\<sigma> n ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3044
        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
  3045
      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
  3046
        using F.MCauchy_def \<sigma> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3047
      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
  3048
        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
  3049
      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
  3050
      proof (rule cSup_upper)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3051
        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
  3052
          using that bd2 by (meson bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3053
      qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3054
      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
  3055
        unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3056
      proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3057
        show "range (\<lambda>n. \<sigma> n x) \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3058
          using \<sigma>M that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3059
        fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3060
        assume "\<epsilon> > 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3061
        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
  3062
          using \<sigma> by (force simp: F.MCauchy_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3063
        { fix n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3064
          assume n: "N \<le> n" "N \<le> n'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3065
          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
  3066
            using that sup by presburger
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3067
          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
  3068
            by (simp add: fdist_def \<open>S \<noteq> {}\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3069
          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
  3070
            by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3071
        } 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
  3072
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3073
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3074
      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
  3075
        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
  3076
      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
  3077
        by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3078
      define g where "g \<equiv> restrict g0 S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3079
      have gext: "g \<in> extensional S" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3080
       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
  3081
        by (auto simp: g_def g0)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3082
      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
  3083
        using glim limitin_metric that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3084
      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
  3085
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3086
        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
  3087
          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
  3088
          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
  3089
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3090
        proof (intro exI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3091
          fix x n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3092
          assume "x \<in> S" and "N \<le> n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3093
          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
  3094
            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
  3095
          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
  3096
            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
  3097
          also have "\<dots> < \<epsilon>/2 + \<epsilon>/2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3098
            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
  3099
          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
  3100
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3101
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3102
      have "limitin F.mtopology \<sigma> g sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3103
        unfolding F.limit_metric_sequentially
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3104
      proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3105
        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
  3106
          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
  3107
        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
  3108
          using \<sigma>M by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3109
        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
  3110
          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
  3111
        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
  3112
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3113
          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
  3114
            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
  3115
          also have "\<dots> \<le> B+1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3116
          proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3117
            have "d a (\<sigma> N x) \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3118
              by (simp add: B that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3119
            moreover 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3120
            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
  3121
            proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3122
              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
  3123
                using 1 dense by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3124
              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
  3125
                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
  3126
              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
  3127
                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
  3128
              also have "\<dots> < 1 + (r-1)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3129
                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
  3130
              finally have "d (\<sigma> N x) (g x) < r"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3131
                by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3132
              with r show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3133
                by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3134
            qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3135
            ultimately show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3136
              by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3137
          qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3138
          finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3139
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3140
        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
  3141
          unfolding mbounded by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3142
        with gwd gext show "g \<in> fspace S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3143
          by (auto simp: fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3144
        fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3145
        assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3146
        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
  3147
          by (meson unif half_gt_zero)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3148
        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
  3149
          using \<open>g \<in> fspace S\<close> False that
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3150
          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
  3151
        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
  3152
          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
  3153
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3154
      then show "\<exists>x. limitin F.mtopology \<sigma> x sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3155
        by blast 
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
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3162
subsection\<open>Metric space of continuous bounded functions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3163
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3164
definition cfunspace where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3165
  "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
  3166
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3167
lemma mspace_cfunspace [simp]:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3168
  "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
  3169
     {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
  3170
         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
  3171
         continuous_map X (mtopology_of m) f}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3172
  by (auto simp: cfunspace_def Metric_space.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3173
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3174
lemma mdist_cfunspace_eq_mdist_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3175
  "mdist (cfunspace X m) = mdist (funspace (topspace X) m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3176
  by (auto simp: cfunspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3177
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3178
lemma cfunspace_subset_funspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3179
   "mspace (cfunspace X m) \<subseteq> mspace (funspace (topspace X) m)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3180
  by (simp add: cfunspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3181
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3182
lemma cfunspace_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3183
   "\<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
  3184
     \<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
  3185
  by (simp add: cfunspace_def Metric_space.funspace_mdist_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3186
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3187
lemma cfunspace_imp_bounded2:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3188
  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
  3189
  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
  3190
  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
  3191
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3192
lemma cfunspace_mdist_lt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3193
   "\<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
  3194
     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
  3195
     x \<in> topspace X\<rbrakk>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3196
     \<Longrightarrow> mdist m (f x) (g x) < a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3197
  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
  3198
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3199
lemma mdist_cfunspace_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3200
  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
  3201
  shows "mdist (cfunspace X m) f g \<le> B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3202
proof (cases "topspace X = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3203
  case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3204
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3205
    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
  3206
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3207
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3208
  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
  3209
    by (meson B bdd_above.I2)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3210
  with assms bdd show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3211
    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
  3212
qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3215
lemma mdist_cfunspace_imp_mdist_le:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3216
   "\<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
  3217
     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
  3218
  using cfunspace_mdist_le by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3219
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3220
lemma compactin_mspace_cfunspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3221
   "compactin X (topspace X)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3222
     \<Longrightarrow> mspace (cfunspace X m) =
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3223
          {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
  3224
               f \<in> extensional (topspace X) \<and>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3225
               continuous_map X (mtopology_of m) f}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3226
  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
  3227
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3228
lemma (in Metric_space) mcomplete_cfunspace:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3229
  assumes "mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3230
  shows "mcomplete_of (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3231
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3232
  interpret F: Metric_space "fspace (topspace X)" "fdist (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3233
    by (simp add: Metric_space_funspace)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3234
  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
  3235
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3236
    show "mspace (cfunspace X Self) \<subseteq> fspace (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3237
      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
  3238
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3239
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3240
  proof (cases "topspace X = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3241
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3242
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3243
      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
  3244
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3245
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3246
    have *: "continuous_map X mtopology g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3247
      if "range \<sigma> \<subseteq> mspace (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3248
        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
  3249
      unfolding continuous_map_to_metric
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3250
    proof (intro strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3251
      have \<sigma>: "\<And>n. continuous_map X mtopology (\<sigma> n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3252
        using that by (auto simp: mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3253
      fix x and \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3254
      assume "x \<in> topspace X" and "0 < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3255
      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
  3256
        unfolding mtopology_of_def F.limitin_metric
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3257
        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
  3258
      then obtain U where "openin X U" "x \<in> U" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3259
        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
  3260
        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
  3261
      moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3262
      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
  3263
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3264
        have "U \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3265
          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
  3266
        have gx: "g x \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3267
          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
  3268
        have "y \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3269
          using \<open>U \<subseteq> topspace X\<close> that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3270
        have gy: "g y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3271
          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
  3272
        have "d (g x) (g y) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3273
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3274
          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
  3275
          proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3276
            have "bdd_above ((\<lambda>x. d (\<sigma> N x) (g x)) ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3277
              by (metis F.limit_metric_sequentially False N bdd_above_dist g order_refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3278
            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
  3279
              by (simp add: cSup_upper that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3280
            also have "\<dots> \<le> \<epsilon>/3"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3281
              by (smt (verit) F.limit_metric_sequentially False N Sup.SUP_cong fdist_def g order_refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3282
            finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3283
          qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3284
          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
  3285
            using U gx gy that triangle by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3286
          also have "\<dots> < \<epsilon>/3 + \<epsilon>/3 + \<epsilon>/3"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3287
            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
  3288
          finally show ?thesis by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3289
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3290
        with gx gy 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
      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
  3293
        by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3294
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3295
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3296
    have "S.sub.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3297
    proof (rule S.sequentially_closedin_mcomplete_imp_mcomplete)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3298
      show "F.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3299
        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
  3300
      fix \<sigma> g
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3301
      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
  3302
      show "g \<in> mspace (cfunspace X Self)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3303
      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
  3304
        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
  3305
          using g F.limitin_mspace by (force simp: fspace_def)+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3306
        show "continuous_map X mtopology g"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3307
          using * g by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3308
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3309
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3310
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3311
      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
  3312
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3313
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3314
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3315
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3316
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
  3317
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3318
lemma (in Metric_space) metric_completion_explicit:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3319
  obtains f :: "['a,'a] \<Rightarrow> real" and S where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3320
      "S \<subseteq> mspace(funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3321
      "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
  3322
      "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3323
      "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
  3324
      "\<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
  3325
            \<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
  3326
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3327
  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
  3328
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3329
  proof (cases "M = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3330
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3331
    then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3332
      using that by (simp add: mcomplete_of_def mcomplete_trivial)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3333
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3334
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3335
    then obtain a where "a \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3336
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3337
    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
  3338
    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
  3339
    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
  3340
      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
  3341
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3342
    have fim: "f ` M \<subseteq> mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3343
    proof (clarsimp simp: m'_def Met_TC.fspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3344
      fix b
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3345
      assume "b \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3346
      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
  3347
        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
  3348
      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
  3349
        by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3350
      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
  3351
        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
  3352
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3353
    show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3354
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3355
      show "S \<subseteq> mspace (funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3356
        by (simp add: S_def in_closure_of subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3357
      have "closedin S.mtopology (S \<inter> Met_TC.fspace M)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3358
        by (simp add: S_def closedin_Int funspace_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3359
      moreover have "S.mcomplete"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3360
        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
  3361
      ultimately show "mcomplete_of (submetric (funspace M euclidean_metric) S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3362
        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
  3363
      show "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3364
        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
  3365
      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
  3366
        by (auto simp: f_def S_def mtopology_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3367
      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
  3368
        if "x \<in> M" "y \<in> M" for x y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3369
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3370
        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
  3371
          using that by (auto simp: f_def dist_real_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3372
        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
  3373
          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
  3374
          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
  3375
          by (smt (verit, best) commute triangle')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3376
        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
  3377
          by (intro cSup_unique) auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3378
        with that fim show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3379
          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
  3380
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3381
    qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3385
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3386
lemma (in Metric_space) metric_completion:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3387
  obtains f :: "['a,'a] \<Rightarrow> real" and m' where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3388
    "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
  3389
    "f \<in> M \<rightarrow> mspace m' "
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3390
    "mtopology_of m' closure_of f ` M = mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3391
    "\<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
  3392
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3393
  obtain f :: "['a,'a] \<Rightarrow> real" and S where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3394
    Ssub: "S \<subseteq> mspace(funspace M euclidean_metric)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3395
    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
  3396
    and fim: "f \<in> M \<rightarrow> S"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3397
    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
  3398
    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
  3399
    using metric_completion_explicit by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3400
  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
  3401
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3402
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3403
    show "mcomplete_of m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3404
      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
  3405
    show "f \<in> M \<rightarrow> mspace m'"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3406
      using Ssub fim m'_def by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3407
    show "mtopology_of m' closure_of f ` M = mspace m'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3408
      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
  3409
      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
  3410
    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
  3411
      using that eqd m'_def by force 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3412
  qed 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3413
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3414
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3415
lemma metrizable_space_completion:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3416
  assumes "metrizable_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3417
  obtains f :: "['a,'a] \<Rightarrow> real" and Y where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3418
       "completely_metrizable_space Y" "embedding_map X Y f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3419
                "Y closure_of (f ` (topspace X)) = topspace Y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3420
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3421
  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
  3422
    using assms metrizable_space_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3423
  then interpret Metric_space M d by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3424
  obtain f :: "['a,'a] \<Rightarrow> real" and m' where
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3425
    "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
  3426
    and fim: "f \<in> M \<rightarrow> mspace m' "
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3427
    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
  3428
    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
  3429
    by (metis metric_completion)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3430
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3431
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3432
    show "completely_metrizable_space (mtopology_of m')"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3433
      using \<open>mcomplete_of m'\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3434
      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
  3435
      by (metis Metric_space_mspace_mdist)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3436
    show "embedding_map X (mtopology_of m') f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3437
      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
  3438
      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
  3439
             mtopology_of_def)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3440
    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
  3441
      by (simp add: Xeq m')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3442
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3443
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3444
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3445
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
subsection\<open>Contractions\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3448
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3449
lemma (in Metric_space) contraction_imp_unique_fixpoint:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3450
  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
  3451
    and "f \<in> M \<rightarrow> M"
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3452
    and "k < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3453
    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
  3454
    and "x \<in> M" "y \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3455
  shows "x = y"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3456
  by (smt (verit, ccfv_SIG) Metric_space.mdist_pos_less Metric_space_axioms assms mult_le_cancel_right1)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3457
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3458
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
  3459
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
  3460
  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
  3461
    and "k < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3462
    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
  3463
  obtains x where "x \<in> M" "f x = x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3464
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3465
  obtain a where "a \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3466
    using \<open>M \<noteq> {}\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3467
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3468
  proof (cases "\<forall>x \<in> M. f x = f a")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3469
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3470
    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
  3471
      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
  3472
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3473
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3474
    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
  3475
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3476
    have "k>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3477
      using Lipschitz_coefficient_pos [where f=f]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3478
      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
  3479
    define \<sigma> where "\<sigma> \<equiv> \<lambda>n. (f^^n) a"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3480
    have f_iter: "\<sigma> n \<in> M" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3481
      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
  3482
    show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3483
    proof (cases "f a = a")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3484
      case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3485
      then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3486
        using \<open>a \<in> M\<close> that by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3487
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3488
      case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3489
      have "MCauchy \<sigma>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3490
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3491
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3492
          unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3493
        proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3494
          show "range \<sigma> \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3495
            using f_iter by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3496
          fix \<epsilon>::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3497
          assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3498
          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
  3499
            by (fastforce simp: divide_simps Pi_iff)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3500
          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
  3501
            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
  3502
          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
  3503
            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
  3504
          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
  3505
          proof (intro exI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3506
            fix n n'
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3507
            assume "n<n'" "N \<le> n" "N \<le> n'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3508
            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
  3509
            proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3510
              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
  3511
              proof (induction m)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3512
                case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3513
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3514
                  by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3515
              next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3516
                case (Suc m)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3517
                then consider "n<m" | "m=n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3518
                  by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3519
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3520
                proof cases
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3521
                  case 1
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3522
                  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
  3523
                    by (simp add: f_iter triangle)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3524
                  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
  3525
                    using Suc 1 by linarith
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3526
                  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
  3527
                    using "1" by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3528
                  finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3529
                qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3530
              qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3531
              with \<open>n < n'\<close> show ?thesis by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3532
            qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3533
            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
  3534
            proof (rule sum_mono)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3535
              fix i
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3536
              assume "i \<in> {n..<n'}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3537
              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
  3538
              proof (induction i)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3539
                case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3540
                then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3541
                  by (auto simp: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3542
              next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3543
                case (Suc i)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3544
                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
  3545
                  using con \<sigma>_def f_iter fim by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3546
                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
  3547
                  using Suc \<open>0 < k\<close> by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3548
                finally show ?case .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3549
              qed
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
            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
  3552
              by (simp add: sum_distrib_left)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3553
            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
  3554
              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
  3555
            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
  3556
              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
  3557
            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
  3558
              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
  3559
              by (simp add: algebra_simps flip: power_add)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3560
            also have "\<dots> < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3561
              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
  3562
              apply (simp add: field_simps)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3563
              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
  3564
            finally show "d (\<sigma> n) (\<sigma> n') < \<epsilon>" .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3565
          qed 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3566
          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
  3567
            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
  3568
        qed
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
      then obtain l where l: "limitin mtopology \<sigma> l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3571
        using \<open>mcomplete\<close> mcomplete_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3572
      show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3573
      proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3574
        show "l \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3575
          using l limitin_mspace by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3576
        show "f l = l"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3577
        proof (rule limitin_metric_unique)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3578
          have "limitin mtopology (f \<circ> \<sigma>) (f l) sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3579
          proof (rule continuous_map_limit)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3580
            have "Lipschitz_continuous_map Self Self f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3581
              using con by (auto simp: Lipschitz_continuous_map_def fim)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3582
            then show "continuous_map mtopology mtopology f"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3583
              using Lipschitz_continuous_imp_continuous_map Self_def by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3584
          qed (use l in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3585
          moreover have "(f \<circ> \<sigma>) = (\<lambda>i. \<sigma>(i+1))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3586
            by (auto simp: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3587
          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
  3588
            using limitin_sequentially_offset_rev [of mtopology \<sigma> 1]
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3589
            by (simp add: \<sigma>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3590
        qed (use l in \<open>auto simp: \<sigma>_def\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3591
      qed
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
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
subsection\<open> The Baire Category Theorem\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3598
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3599
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
  3600
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3601
lemma (in Metric_space) metric_Baire_category:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3602
  assumes "mcomplete" "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3603
  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
  3604
shows "mtopology closure_of \<Inter>\<G> = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3605
proof (cases "\<G>={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3606
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3607
  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
  3608
    by (metis \<open>countable \<G>\<close> uncountable_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3609
  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
  3610
    by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3611
  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
  3612
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3613
    have "W \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3614
      using openin_mtopology W by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3615
    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
  3616
      if "r>0" "x \<in> M" for x r n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3617
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3618
      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
  3619
        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
  3620
        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
  3621
      then have "z \<in> M" by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3622
      have "openin mtopology (U n \<inter> mball x r)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3623
        by (simp add: openin_Int u_open)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3624
      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
  3625
        by (meson openin_mtopology_mcball)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3626
      define r' where "r' \<equiv> min e (r/4)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3627
      show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3628
      proof (intro exI conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3629
        show "0 < r'" "r' < r / 2" "z \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3630
          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
  3631
        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
  3632
          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
  3633
      qed
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
    then obtain nextr nextx 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3636
      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
  3637
        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
  3638
        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
  3639
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3640
    obtain x0 where x0: "x0 \<in> U 0 \<inter> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3641
      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
  3642
    then have "x0 \<in> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3643
      using \<open>W \<subseteq> M\<close> by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3644
    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
  3645
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3646
      have "openin mtopology (U 0 \<inter> W)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3647
        using W u_open by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3648
      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
  3649
        by (meson Int_subset_iff openin_mtopology x0)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3650
      define r0 where "r0 \<equiv> (min r 1) / 2"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3651
      show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3652
      proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3653
        show "0 < r0" "r0 < 1"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3654
          using \<open>r>0\<close> by (auto simp: r0_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3655
        show "mcball x0 r0 \<subseteq> U 0 \<inter> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3656
          using r \<open>0 < r0\<close> r0_def by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3657
      qed
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
    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
  3660
    have b0[simp]: "b 0 = (x0,r0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3661
      by (simp add: b_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3662
    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
  3663
      by (simp add: b_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3664
    define xf where "xf \<equiv> fst \<circ> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3665
    define rf where "rf \<equiv> snd \<circ> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3666
    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
  3667
    proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3668
      case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3669
      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
  3670
        by (auto simp: rf_def xf_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3671
    next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3672
      case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3673
      then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3674
        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
  3675
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3676
    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
  3677
      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
  3678
    have half: "rf (Suc n) < rf n / 2" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3679
      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
  3680
    then have "decseq rf"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3681
      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
  3682
    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
  3683
      using that
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3684
    proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3685
      case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3686
      then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3687
        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
  3688
    qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3689
    have "MCauchy xf"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3690
      unfolding MCauchy_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3691
    proof (intro conjI strip)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3692
      show "range xf \<subseteq> M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3693
        using rfxf by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3694
      fix \<epsilon> :: real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3695
      assume "\<epsilon>>0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3696
      then obtain N where N: "inverse (2^N) < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3697
        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
  3698
      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
  3699
      proof -           
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3700
        have *: "rf n < inverse (2 ^ n)" for n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3701
        proof (induction n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3702
          case 0
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3703
          then show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3704
            by (simp add: \<open>r0 < 1\<close> rf_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3705
        next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3706
          case (Suc n)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3707
          with half show ?case
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3708
            by simp (smt (verit))
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3709
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3710
        have "rf n \<le> rf N"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3711
          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
  3712
        moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3713
        have "xf n' \<in> mball (xf n) (rf n)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3714
          using nested rfxf \<open>n \<le> n'\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3715
        ultimately have "d (xf n) (xf n') < rf N"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3716
          by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3717
        also have "\<dots> < \<epsilon>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3718
          using "*" N order.strict_trans by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3719
        finally show ?thesis .
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3720
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3721
      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
  3722
        by (metis commute linorder_le_cases)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3723
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3724
    then obtain l where l: "limitin mtopology xf l sequentially"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3725
      using \<open>mcomplete\<close> mcomplete_alt by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3726
    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
  3727
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3728
      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
  3729
        unfolding eventually_sequentially
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3730
        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
  3731
      with l limitin_closedin show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3732
        by (metis closedin_mcball trivial_limit_sequentially)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3733
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3734
    then have "\<And>n. l \<in> U n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3735
      using mcball_sub by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3736
    moreover have "l \<in> W"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3737
      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
  3738
    ultimately show ?thesis by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3739
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3740
  with U show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3741
    by (metis dense_intersects_open topspace_mtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3742
qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3743
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
lemma (in Metric_space) metric_Baire_category_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3747
  assumes "mcomplete" "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3748
    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
  3749
  shows "mtopology interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3750
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3751
  have *: "mtopology closure_of \<Inter>((-)M ` \<G>) = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3752
  proof (intro metric_Baire_category conjI \<open>mcomplete\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3753
    show "countable ((-) M ` \<G>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3754
      using \<open>countable \<G>\<close> by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3755
    fix T
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3756
    assume "T \<in> (-) M ` \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3757
    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
  3758
      using empty metric_closedin_iff_sequentially_closed by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3759
    with empty show "openin mtopology T" by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3760
    show "mtopology closure_of T = M"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3761
      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
  3762
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3763
  with closure_of_eq show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3764
    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
  3765
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3766
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
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
  3769
lemma Baire_category_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3770
  assumes "locally_compact_space X" "regular_space X" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3771
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3772
    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
  3773
  shows "X interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3774
proof (cases "\<G> = {}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3775
  case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3776
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3777
    by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3778
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3779
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3780
  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
  3781
    by (metis \<open>countable \<G>\<close> uncountable_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3782
  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
  3783
    by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3784
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3785
  proof (clarsimp simp: T interior_of_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3786
    fix z U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3787
    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
  3788
    with openin_subset have "z \<in> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3789
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3790
    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
  3791
      using assms locally_compact_regular_space_neighbourhood_base by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3792
    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
  3793
      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
  3794
    have nb_closedin: "neighbourhood_base_of (closedin X) X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3795
      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
  3796
    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
  3797
        \<Phi> (Suc n) \<subseteq> \<Phi> n"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3798
    proof (rule dependent_nat_choice)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3799
      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
  3800
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3801
        have False if "V \<subseteq> T 0"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3802
          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
  3803
        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
  3804
          using T \<open>openin X V\<close> empty by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3805
        with nb_closedin 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3806
        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
  3807
          unfolding neighbourhood_base_of by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3808
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3809
        proof (intro exI conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3810
          show "C \<subseteq> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3811
            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
  3812
          show "X interior_of C \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3813
            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
  3814
          show "disjnt C (T 0)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3815
            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
  3816
        qed (use \<open>closedin X C\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3817
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3818
      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
  3819
        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
  3820
        for C n
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3821
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3822
        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
  3823
          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
  3824
        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
  3825
          using T empty by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3826
        with nb_closedin 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3827
        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
  3828
          unfolding neighbourhood_base_of by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3829
        show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3830
        proof (intro conjI exI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3831
          show "D \<subseteq> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3832
            using D interior_of_subset \<section> by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3833
          show "X interior_of D \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3834
            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
  3835
          show "disjnt D (T (Suc n))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3836
            using D disjnt_iff by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3837
          show "D \<subseteq> C"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3838
            using interior_of_subset [of X C] D by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3839
        qed (use \<open>closedin X D\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3840
      qed
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
    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
  3843
      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
  3844
    then have "decseq \<Phi>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3845
      by (simp add: decseq_SucI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3846
    moreover have "\<And>n. \<Phi> n \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3847
      by (metis \<Phi> bot.extremum_uniqueI interior_of_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3848
    ultimately have "\<Inter>(range \<Phi>) \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3849
      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
  3850
    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
  3851
      using Asub by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3852
    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
  3853
      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
  3854
    ultimately show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3855
      by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3856
  qed
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
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
lemma Baire_category_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3861
  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
  3862
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3863
    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
  3864
  shows "X interior_of \<Union>\<G> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3865
  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
  3866
  by (metis assms completely_metrizable_space_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3867
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
lemma Baire_category:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3870
  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
  3871
    and "countable \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3872
    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
  3873
  shows "X closure_of \<Inter>\<G> = topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3874
proof (cases "\<G>={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3875
  case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3876
  have *: "X interior_of \<Union>((-)(topspace X) ` \<G>) = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3877
    proof (intro Baire_category_alt conjI assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3878
  show "countable ((-) (topspace X) ` \<G>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3879
    using assms by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3880
    fix T
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3881
    assume "T \<in> (-) (topspace X) ` \<G>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3882
    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
  3883
      by (meson top image_iff openin_subset)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3884
    then show "closedin X T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3885
      by (simp add: closedin_diff top)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3886
    show "X interior_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3887
      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
  3888
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3889
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3890
    by (simp add: closure_of_eq_topspace interior_of_complement)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3891
qed auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3892
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
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
  3895
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3896
lemma locally_connected_not_countable_closed_union:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3897
  assumes "topspace X \<noteq> {}" and csX: "connected_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3898
    and lcX: "locally_connected_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3899
    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
  3900
    and "countable \<U>" and pwU: "pairwise disjnt \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3901
    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
  3902
    and UU_eq: "\<Union>\<U> = topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3903
  shows "\<U> = {topspace X}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3904
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3905
  define \<V> where "\<V> \<equiv> (frontier_of) X ` \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3906
  define B where "B \<equiv> \<Union>\<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3907
  then have Bsub: "B \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3908
    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
  3909
  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
  3910
    by (meson clo closedin_def that)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3911
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3912
  proof (rule ccontr)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3913
    assume "\<U> \<noteq> {topspace X}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3914
    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
  3915
      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
  3916
    then have "B \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3917
      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
  3918
    moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3919
    have "subtopology X B interior_of B = B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3920
      by (simp add: Bsub interior_of_openin openin_subtopology_refl)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3921
    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
  3922
      by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3923
    have "subtopology X B interior_of \<Union>\<V> = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3924
    proof (intro Baire_category_alt conjI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3925
      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
  3926
        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
  3927
      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
  3928
        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
  3929
      moreover have "disjnt B (\<Union>((interior_of) X ` \<U>))"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3930
        using pwU 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3931
        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
  3932
        by (metis clo closure_of_eq interior_of_subset subsetD)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3933
      ultimately have "B = topspace X - \<Union> ((interior_of) X ` \<U>)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3934
        by (auto simp: UU_eq disjnt_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3935
      then have "closedin X B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3936
        by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3937
      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
  3938
        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
  3939
            locally_compact_space_closed_subset regular_space_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3940
      show "countable \<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3941
        by (simp add: \<V>_def \<open>countable \<U>\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3942
      fix V
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3943
      assume "V \<in> \<V>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3944
      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
  3945
        by (auto simp: \<V>_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3946
      show "closedin (subtopology X B) V"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3947
        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
  3948
      have "subtopology X B interior_of (X frontier_of S) = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3949
      proof (clarsimp simp: interior_of_def openin_subtopology_alt)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3950
        fix a U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3951
        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
  3952
        then have "a \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3953
          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
  3954
        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
  3955
          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
  3956
        have "W \<inter> X frontier_of S \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3957
          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
  3958
        with frontier_of_openin_straddle_Int
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3959
        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
  3960
          using \<open>openin X W\<close> by (metis openin_subset) 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3961
        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
  3962
          by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3963
        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
  3964
          by auto 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3965
        then have "disjnt S T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3966
          by (metis \<open>S \<in> \<U>\<close> pairwise_def pwU)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3967
        then have "C - T \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3968
          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
  3969
        then have "C \<inter> X frontier_of T \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3970
          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
  3971
        moreover have "C \<inter> X frontier_of T = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3972
        proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3973
          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
  3974
            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
  3975
          moreover have "X frontier_of T \<union> B = B"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3976
            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
  3977
          ultimately show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3978
            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
  3979
        qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3980
        ultimately show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3981
          by simp
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3982
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3983
      with S show "subtopology X B interior_of V = {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3984
        by meson
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3985
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3986
    then show False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3987
      using B_def int_B_nonempty by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3988
  qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3991
lemma real_Sierpinski_lemma:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3992
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3993
  assumes "a \<le> b"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3994
    and "countable \<U>" and pwU: "pairwise disjnt \<U>"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3995
    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
  3996
    and "\<Union>\<U> = {a..b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3997
  shows "\<U> = {{a..b}}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3998
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  3999
  have "locally_connected_space (top_of_set {a..b})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4000
    by (simp add: locally_connected_real_interval)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4001
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4002
  have "completely_metrizable_space (top_of_set {a..b})"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4003
    by (metis box_real(2) completely_metrizable_space_cbox)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4004
  ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4005
  show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4006
    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
  4007
    apply (simp add: closedin_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4008
    by (metis Union_upper inf.orderE)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4009
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4010
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
subsection\<open>The Tychonoff embedding\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4013
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4014
lemma completely_regular_space_cube_embedding_explicit:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4015
  assumes "completely_regular_space X" "Hausdorff_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4016
  shows "embedding_map X
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4017
             (product_topology (\<lambda>f. top_of_set {0..1::real})
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4018
                (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
  4019
                  {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
  4020
             (\<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
  4021
              f x)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4022
proof -
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78202
diff changeset
  4023
  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
  4024
  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
  4025
  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
  4026
  proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4027
    have "closedin X {x}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4028
      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
  4029
    then obtain f where contf: "continuous_map X euclideanreal f" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4030
      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
  4031
      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
  4032
      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
  4033
    then have "bounded (f ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4034
      by (meson bounded_closed_interval bounded_subset image_subset_iff)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4035
    with contf f01 have "restrict f (topspace X) \<in> K"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4036
      by (auto simp: K_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4037
    with fxy xy show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4038
      unfolding e_def by (metis restrict_apply' zero_neq_one)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4039
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4040
  then have "inj_on e (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4041
    by (meson inj_onI)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4042
  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
  4043
    by (metis inv_into_f_f)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4044
  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
  4045
  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
  4046
    fix x U
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4047
    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
  4048
    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
  4049
          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
  4050
      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
  4051
      by (metis Diff_iff openin_closedin_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4052
    then have "bounded (g ` topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4053
      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
  4054
    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
  4055
      using contg by (simp add: continuous_map_def)
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4056
    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
  4057
      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
  4058
    have "openin (top_of_set {0..1}) {0..<1::real}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4059
      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
  4060
    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
  4061
      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
  4062
    moreover have "e y = e x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4063
      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
  4064
           and y: "y \<in> topspace X" for y
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4065
    proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4066
      have "e y (restrict g (topspace X)) \<in> {0..<1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4067
        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
  4068
    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
  4069
      by (fastforce simp: e_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4070
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4071
    ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4072
    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
  4073
      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
  4074
      by (auto simp: openin_PiE_gen e')
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4075
  qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4076
  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
  4077
    unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4078
    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
  4079
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4080
    by (simp add: K_def e_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4081
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4082
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
lemma completely_regular_space_cube_embedding:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4085
  fixes X :: "'a topology"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4086
  assumes "completely_regular_space X" "Hausdorff_space X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4087
  obtains K:: "('a\<Rightarrow>real)set" and e
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4088
    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
  4089
  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
  4090
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
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
  4093
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4094
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
  4095
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
  4096
we factor through the Kolmogorov quotient." -- John Harrison\<close>
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4097
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4098
lemma Urysohn_completely_regular_closed_compact:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4099
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4100
  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
  4101
  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
  4102
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4103
  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
  4104
    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
  4105
  proof (cases "T={}")
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4106
    case True
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4107
    show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4108
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4109
      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
  4110
        using True by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4111
    qed   
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4112
  next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4113
    case False
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4114
    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
  4115
      using assms unfolding completely_regular_space_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4116
      by (meson DiffI compactin_subset_topspace disjnt_iff subset_eq)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4117
    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
  4118
                  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
  4119
                  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
  4120
      by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4121
    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
  4122
      by (meson continuous_map_in_subtopology)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4123
    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
  4124
    have "Ball (G`T) (openin X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4125
      using contg unfolding G_def continuous_map_in_subtopology
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4126
      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
  4127
    moreover have "T \<subseteq> \<Union>(G`T)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4128
      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
  4129
    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
  4130
      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
  4131
    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
  4132
      by (metis finite_subset_image)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4133
    with False have "K \<noteq> {}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4134
      by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4135
    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
  4136
    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
  4137
      by force
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4138
    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
  4139
      by (simp add: max_def_raw)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4140
    show thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4141
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4142
      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
  4143
        using \<open>K \<subseteq> T\<close> g1 that by auto
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4144
      then show "f ` S \<subseteq> {1}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4145
        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
  4146
      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
  4147
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4148
        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
  4149
          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
  4150
        then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4151
          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
  4152
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4153
      then show "f ` T \<subseteq> {0}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4154
        by (force simp: f_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4155
      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
  4156
        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
  4157
      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
  4158
      proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4159
        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
  4160
          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
  4161
        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
  4162
          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
  4163
        then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4164
          by (simp add: max_def_raw)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4165
      qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4166
      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
  4167
        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
  4168
    qed
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
  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
  4171
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4172
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4173
    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
  4174
      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
  4175
      by (smt (verit, best) mult_right_le_one_le)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4176
    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
  4177
      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
  4178
      by (intro conjI continuous_intros; simp)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4179
    show "g ` T \<subseteq> {a}" "g ` S \<subseteq> {b}"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4180
      using f0 f1 by (auto simp: g_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4181
  qed 
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
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
lemma Urysohn_completely_regular_compact_closed:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4186
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4187
  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
  4188
  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
  4189
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4190
  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
  4191
    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
  4192
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4193
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4194
    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
  4195
      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
  4196
    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
  4197
      using fim by fastforce+
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4198
  qed
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
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4201
lemma Urysohn_completely_regular_compact_closed_alt:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4202
  fixes a b::real
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4203
  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
  4204
  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
  4205
proof (cases a b rule: le_cases)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4206
  case le
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4207
  then show ?thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4208
    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
  4209
next
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4210
  case ge
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4211
  then show ?thesis 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4212
    using Urysohn_completely_regular_closed_compact assms
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4213
    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
  4214
qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4215
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
lemma Tietze_extension_comp_reg_aux:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4218
  fixes T :: "real set"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4219
  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
  4220
    and T: "is_interval T" "T\<noteq>{}" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4221
    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
  4222
  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
  4223
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4224
  obtain K:: "('a\<Rightarrow>real)set" and e
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4225
    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
  4226
    using assms completely_regular_space_cube_embedding by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4227
  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
  4228
  have e: "embedding_map X cube e"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4229
    using e0 by (simp add: cube_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4230
  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
  4231
    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
  4232
  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
  4233
     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
  4234
     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
  4235
    by (auto simp: homeomorphic_maps_def)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4236
  have "Hausdorff_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4237
    unfolding cube_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4238
    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
  4239
  have "normal_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4240
  proof (rule compact_Hausdorff_or_regular_imp_normal_space)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4241
    show "compact_space cube"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4242
      unfolding cube_def
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4243
      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
  4244
  qed (use \<open>Hausdorff_space cube\<close> in auto)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4245
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4246
  have comp: "compactin cube (e ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4247
    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
  4248
  then have "closedin cube (e ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4249
    by (intro compactin_imp_closedin \<open>Hausdorff_space cube\<close>)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4250
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4251
  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
  4252
  proof (intro continuous_map_compose)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4253
    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
  4254
      unfolding continuous_map_in_subtopology
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4255
    proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4256
      show "continuous_map (subtopology cube (e ` S)) X e'"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4257
        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
  4258
      show "e' ` topspace (subtopology cube (e ` S)) \<subseteq> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4259
        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
  4260
    qed
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4261
  qed (simp add: contf)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4262
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4263
  have "(f \<circ> e') ` e ` S \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4264
    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
  4265
  ultimately
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4266
  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
  4267
                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
  4268
    using Tietze_extension_realinterval T by metis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4269
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4270
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4271
    show "continuous_map X euclideanreal (g \<circ> e)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4272
      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
  4273
    show "(g \<circ> e) ` topspace X \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4274
      using gsub conte continuous_map_image_subset_topspace by fastforce
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4275
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4276
    assume "x \<in> S"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4277
    then show "(g \<circ> e) x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4278
      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
  4279
  qed
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
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
lemma Tietze_extension_completely_regular:
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4284
  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
  4285
    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
  4286
  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
  4287
    "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4288
proof -
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4289
  define Q where "Q \<equiv> Kolmogorov_quotient X ` (topspace X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4290
  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
  4291
    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
  4292
    using Kolmogorov_quotient_lift_exists 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4293
    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
  4294
  have "S \<subseteq> topspace X"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4295
    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
  4296
  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
  4297
    using Q_def by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4298
  have creg: "completely_regular_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4299
    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
  4300
  then have "regular_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4301
    by (simp add: completely_regular_imp_regular_space)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4302
  then have "Hausdorff_space (subtopology X Q)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4303
    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
  4304
  moreover
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4305
  have "compactin (subtopology X Q) (Kolmogorov_quotient X ` S)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4306
    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
  4307
  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
  4308
              and him: "h ` topspace (subtopology X Q) \<subseteq> T" 
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4309
              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
  4310
    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
  4311
    apply (simp add: subtopology_subtopology creg contg assms)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4312
    using fim gf by blast
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4313
  show thesis
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4314
  proof
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4315
    show "continuous_map X euclideanreal (h \<circ> Kolmogorov_quotient X)"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4316
      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
  4317
    show "(h \<circ> Kolmogorov_quotient X) ` topspace X \<subseteq> T"
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4318
      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
  4319
    fix x
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4320
    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
  4321
      by (simp add: gf hg)
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  4322
  qed
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
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4325
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
  4326
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4327
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
  4328
  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
  4329
  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
  4330
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4331
  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
  4332
    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
  4333
  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
  4334
    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
  4335
  have "(UNIV::real set) \<lesssim> {0..1::real}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4336
    using card_eq_real_subset 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4337
    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
  4338
  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
  4339
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4340
    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
  4341
      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
  4342
      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
  4343
      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
  4344
      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
  4345
      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
  4346
    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
  4347
    proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4348
      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
  4349
        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
  4350
      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
  4351
        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
  4352
      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
  4353
        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
  4354
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4355
        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
  4356
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4357
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4358
      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
  4359
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4360
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4361
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4362
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
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
  4365
  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
  4366
  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
  4367
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4368
  have "(UNIV::real set) \<lesssim> {0..1::real}"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4369
    by (metis atLeastAtMost_iff card_eq_real_subset eqpoll_imp_lepoll eqpoll_sym less_le_not_le zero_less_one)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4370
  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
  4371
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4372
    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
  4373
       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
  4374
       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
  4375
      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
  4376
    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
  4377
    proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4378
      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
  4379
        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
  4380
      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
  4381
        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
  4382
      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
  4383
        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
  4384
      ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4385
        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
  4386
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4387
    then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4388
      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
  4389
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4390
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4391
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4392
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4393
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
  4394
  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
  4395
  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
  4396
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4397
  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
  4398
    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
  4399
  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
  4400
    by force
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4401
  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
  4402
    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
  4403
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4404
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4405
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
  4406
   "\<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
  4407
  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
  4408
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4409
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
  4410
   "\<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
  4411
  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
  4412
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4413
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
  4414
  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
  4415
  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
  4416
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4417
  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
  4418
    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
  4419
  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
  4420
    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
  4421
  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
  4422
  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
  4423
    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
  4424
      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
  4425
    fix V n
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4426
    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
  4427
    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
  4428
      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
  4429
    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
  4430
      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
  4431
    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
  4432
      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
  4433
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4434
  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
  4435
    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
  4436
  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
  4437
    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
  4438
  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
  4439
    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
  4440
  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
  4441
    by metis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4442
  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
  4443
    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
  4444
  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
  4445
    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
  4446
  ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4447
    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
  4448
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4449
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4450
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
  4451
  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
  4452
  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
  4453
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4454
  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
  4455
    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
  4456
  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
  4457
    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
  4458
  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
  4459
  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
  4460
  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
  4461
    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
  4462
      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
  4463
    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
  4464
      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
  4465
    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
  4466
      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
  4467
    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
  4468
      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
  4469
    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
  4470
      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
  4471
    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
  4472
      by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4473
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4474
  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
  4475
    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
  4476
  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
  4477
    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
  4478
  finally show ?thesis .
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4479
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4480
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4481
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
  4482
  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
  4483
  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
  4484
proof
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4485
  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
  4486
  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
  4487
    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
  4488
  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
  4489
    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
  4490
  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
  4491
    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
  4492
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4493
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4494
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
  4495
  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
  4496
  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
  4497
proof 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4498
  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
  4499
  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
  4500
    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
  4501
  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
  4502
    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
  4503
  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
  4504
    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
  4505
  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
  4506
  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
  4507
  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
  4508
    show "countable \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4509
      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
  4510
    show "disjoint \<A>"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4511
      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
  4512
    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
  4513
      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
  4514
    fix C
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4515
    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
  4516
    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
  4517
      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
  4518
    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
  4519
      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
  4520
    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
  4521
      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
  4522
    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
  4523
      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
  4524
  qed auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4525
  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
  4526
    by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4527
  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
  4528
    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
  4529
  then show False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4530
    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
  4531
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4532
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
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
  4535
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4536
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
  4537
  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
  4538
  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
  4539
         M \<noteq> {} \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4540
         (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
  4541
            \<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
  4542
                           (\<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
  4543
proof (cases "M = {}")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4544
  case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4545
  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
  4546
    by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4547
next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4548
  case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4549
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4550
  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
  4551
    case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4552
    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
  4553
    show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4554
    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
  4555
      case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4556
      have ?R
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4557
        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
  4558
      proof (intro strip)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4559
        fix \<epsilon>::real
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4560
        assume "\<epsilon>>0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4561
        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
  4562
        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
  4563
          and Uless: "\<forall>x\<in>U - {a}. x \<in> S \<longrightarrow> f x \<in> M \<and> d (f x) l < \<epsilon>/2"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4564
          unfolding limitin_metric eventually_within_imp eventually_atin
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4565
          using half_gt_zero by blast
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
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4655
            apply (simp add: eventually_atin eventually_within_imp del: divide_const_simps flip: 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
  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)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4885
    also have "\<dots> homeomorphic_space (prod_topology mtopology (powertop_real {}))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4886
      by (metis PiE_empty_domain homeomorphic_space_sym prod_topology_homeomorphic_space_left topspace_product_topology)
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)})"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  4889
      by (smt (verit) True subtopology_topspace top)
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
78202
759c71cdaf2a More metric space material
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  5185
end