src/HOL/Analysis/T1_Spaces.thy
author wenzelm
Sun, 24 Dec 2023 11:58:33 +0100
changeset 79346 f86c310327df
parent 78336 6bae28577994
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
     1
section\<open>T1 and Hausdorff spaces\<close>
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
theory T1_Spaces
71200
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 70196
diff changeset
     4
imports Product_Topology
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
begin
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
     7
section\<open>T1 spaces with equivalences to many naturally "nice" properties. \<close>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
     8
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
definition t1_space where
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
 "t1_space X \<equiv> \<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x\<noteq>y \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> y \<notin> U)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
lemma t1_space_expansive:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
   "\<lbrakk>topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> t1_space X \<Longrightarrow> t1_space Y"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
  by (metis t1_space_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
lemma t1_space_alt:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
   "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x\<noteq>y \<longrightarrow> (\<exists>U. closedin X U \<and> x \<in> U \<and> y \<notin> U))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
 by (metis DiffE DiffI closedin_def openin_closedin_eq t1_space_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
    20
lemma t1_space_empty [iff]: "t1_space trivial_topology"
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  by (simp add: t1_space_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
lemma t1_space_derived_set_of_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. X derived_set_of {x} = {})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  apply (simp add: t1_space_def derived_set_of_def, safe)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
   apply (metis openin_topspace)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  by force
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
lemma t1_space_derived_set_of_finite:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
   "t1_space X \<longleftrightarrow> (\<forall>S. finite S \<longrightarrow> X derived_set_of S = {})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
proof (intro iffI allI impI)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  fix S :: "'a set"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  assume "finite S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  then have fin: "finite ((\<lambda>x. {x}) ` (topspace X \<inter> S))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
    by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  assume "t1_space X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  then have "X derived_set_of (\<Union>x \<in> topspace X \<inter> S. {x}) = {}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
    unfolding derived_set_of_Union [OF fin]
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
    by (auto simp: t1_space_derived_set_of_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  then have "X derived_set_of (topspace X \<inter> S) = {}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  then show "X derived_set_of S = {}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
qed (auto simp: t1_space_derived_set_of_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
lemma t1_space_closedin_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
   "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. closedin X {x})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  apply (rule iffI)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
  apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  using t1_space_alt by auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
78093
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    52
lemma continuous_closed_imp_proper_map:
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    53
   "\<lbrakk>compact_space X; t1_space Y; continuous_map X Y f; closed_map X Y f\<rbrakk> \<Longrightarrow> proper_map X Y f"
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    54
  unfolding proper_map_def
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    55
  by (smt (verit) closedin_compact_space closedin_continuous_map_preimage 
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    56
      Collect_cong singleton_iff t1_space_closedin_singleton)
cec875dcc59e Finally, the abstract metric space development
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
    57
77943
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77935
diff changeset
    58
lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77935
diff changeset
    59
  by (simp add: t1_space_closedin_singleton)
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77935
diff changeset
    60
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
lemma closedin_t1_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
   "\<lbrakk>t1_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {a}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  by (simp add: t1_space_closedin_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
lemma t1_space_closedin_finite:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
   "t1_space X \<longleftrightarrow> (\<forall>S. finite S \<and> S \<subseteq> topspace X \<longrightarrow> closedin X S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  apply (rule iffI)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  by (simp add: t1_space_closedin_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
lemma closure_of_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
   "t1_space X \<Longrightarrow> X closure_of {a} = (if a \<in> topspace X then {a} else {})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
  by (simp add: closure_of_eq t1_space_closedin_singleton closure_of_eq_empty_gen)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
lemma separated_in_singleton:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  assumes "t1_space X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  shows "separatedin X {a} S \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
        "separatedin X S {a} \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  unfolding separatedin_def
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  using assms closure_of closure_of_singleton by fastforce+
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
lemma t1_space_openin_delete:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
   "t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (U - {x}))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
  apply (rule iffI)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  apply (meson closedin_t1_singleton in_mono openin_diff openin_subset)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
  by (simp add: closedin_def t1_space_closedin_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
lemma t1_space_openin_delete_alt:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
   "t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<longrightarrow> openin X (U - {x}))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  by (metis Diff_empty Diff_insert0 t1_space_openin_delete)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
lemma t1_space_singleton_Inter_open:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
      "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<Inter>{U. openin X U \<and> x \<in> U} = {x})"  (is "?P=?Q")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  and t1_space_Inter_open_supersets:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
     "t1_space X \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<longrightarrow> \<Inter>{U. openin X U \<and> S \<subseteq> U} = S)" (is "?P=?R")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
  have "?R \<Longrightarrow> ?Q"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
    apply clarify
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    apply (drule_tac x="{x}" in spec, simp)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
    done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
  moreover have "?Q \<Longrightarrow> ?P"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
    apply (clarsimp simp add: t1_space_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
    apply (drule_tac x=x in bspec)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
     apply (simp_all add: set_eq_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
    by (metis (no_types, lifting))
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  moreover have "?P \<Longrightarrow> ?R"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  proof (clarsimp simp add: t1_space_closedin_singleton, rule subset_antisym)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
    fix S
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    assume S: "\<forall>x\<in>topspace X. closedin X {x}" "S \<subseteq> topspace X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
    then show "\<Inter> {U. openin X U \<and> S \<subseteq> U} \<subseteq> S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
      apply clarsimp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
      by (metis Diff_insert_absorb Set.set_insert closedin_def openin_topspace subset_insert)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  qed force
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  ultimately show "?P=?Q" "?P=?R"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
lemma t1_space_derived_set_of_infinite_openin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
   "t1_space X \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
        (\<forall>S. X derived_set_of S =
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
             {x \<in> topspace X. \<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite(S \<inter> U)})"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
         (is "_ = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
  assume "t1_space X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  proof safe
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    fix S x U
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    assume "x \<in> X derived_set_of S" "x \<in> U" "openin X U" "finite (S \<inter> U)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
    with \<open>t1_space X\<close> show "False"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
      apply (simp add: t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
      by (metis IntI empty_iff empty_subsetI inf_commute openin_Int_derived_set_of_subset subset_antisym)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
    fix S x
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
    have eq: "(\<exists>y. (y \<noteq> x) \<and> y \<in> S \<and> y \<in> T) \<longleftrightarrow> ~((S \<inter> T) \<subseteq> {x})" for x S T
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
      by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
    assume "x \<in> topspace X" "\<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite (S \<inter> U)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
    then show "x \<in> X derived_set_of S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
      apply (clarsimp simp add: derived_set_of_def eq)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
      by (meson finite.emptyI finite.insertI finite_subset)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  qed (auto simp: in_derived_set_of)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
qed (auto simp: t1_space_derived_set_of_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
lemma finite_t1_space_imp_discrete_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
   "\<lbrakk>topspace X = U; finite U; t1_space X\<rbrakk> \<Longrightarrow> X = discrete_topology U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  by (metis discrete_topology_unique_derived_set t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
lemma t1_space_subtopology: "t1_space X \<Longrightarrow> t1_space(subtopology X U)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  by (simp add: derived_set_of_subtopology t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
lemma closedin_derived_set_of_gen:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
   "t1_space X \<Longrightarrow> closedin X (X derived_set_of S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  apply (clarsimp simp add: in_derived_set_of closedin_contains_derived_set derived_set_of_subset_topspace)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
  by (metis DiffD2 insert_Diff insert_iff t1_space_openin_delete)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
lemma derived_set_of_derived_set_subset_gen:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
   "t1_space X \<Longrightarrow> X derived_set_of (X derived_set_of S) \<subseteq> X derived_set_of S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
  by (meson closedin_contains_derived_set closedin_derived_set_of_gen)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
lemma subtopology_eq_discrete_topology_gen_finite:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
   "\<lbrakk>t1_space X; finite S\<rbrakk> \<Longrightarrow> subtopology X S = discrete_topology(topspace X \<inter> S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
  by (simp add: subtopology_eq_discrete_topology_gen t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
lemma subtopology_eq_discrete_topology_finite:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
   "\<lbrakk>t1_space X; S \<subseteq> topspace X; finite S\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
        \<Longrightarrow> subtopology X S = discrete_topology S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  by (simp add: subtopology_eq_discrete_topology_eq t1_space_derived_set_of_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
lemma t1_space_closed_map_image:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
   "\<lbrakk>closed_map X Y f; f ` (topspace X) = topspace Y; t1_space X\<rbrakk> \<Longrightarrow> t1_space Y"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  by (metis closed_map_def finite_subset_image t1_space_closedin_finite)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
lemma homeomorphic_t1_space: "X homeomorphic_space Y \<Longrightarrow> (t1_space X \<longleftrightarrow> t1_space Y)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  apply (clarsimp simp add: homeomorphic_space_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
  by (meson homeomorphic_eq_everything_map homeomorphic_maps_map t1_space_closed_map_image)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
proposition t1_space_product_topology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
   "t1_space (product_topology X I)
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   179
\<longleftrightarrow> (product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. t1_space (X i))"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   180
proof (cases "(product_topology X I) = trivial_topology")
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  case True
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  then show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   183
    using True t1_space_empty by force
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
  case False
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  then obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   187
    using discrete_topology_unique by (fastforce iff: null_topspace_iff_trivial)
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
  have "t1_space (product_topology X I) \<longleftrightarrow> (\<forall>i\<in>I. t1_space (X i))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
  proof (intro iffI ballI)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
    show "t1_space (X i)" if "t1_space (product_topology X I)" and "i \<in> I" for i
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
    proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
      have clo: "\<And>h. h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<Longrightarrow> closedin (product_topology X I) {h}"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
        using that by (simp add: t1_space_closedin_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
      show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
        unfolding t1_space_closedin_singleton
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
      proof clarify
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
        show "closedin (X i) {xi}" if "xi \<in> topspace (X i)" for xi
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
          using clo [of "\<lambda>j \<in> I. if i=j then xi else f j"] f that \<open>i \<in> I\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
          by (fastforce simp add: closedin_product_topology_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
      qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    show "t1_space (product_topology X I)" if "\<forall>i\<in>I. t1_space (X i)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
      using that
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
      by (simp add: t1_space_closedin_singleton Ball_def PiE_iff closedin_product_topology_singleton)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
  then show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
    using False by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   212
lemma t1_space_prod_topology:
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   213
   "t1_space(prod_topology X Y) \<longleftrightarrow> (prod_topology X Y) = trivial_topology \<or> t1_space X \<and> t1_space Y"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   214
proof (cases "(prod_topology X Y) = trivial_topology")
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   215
  case True then show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   216
  by auto
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   217
next
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   218
  case False
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   219
  have eq: "{(x,y)} = {x} \<times> {y}" for x::'a and y::'b
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   220
    by simp
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   221
  have "t1_space (prod_topology X Y) \<longleftrightarrow> (t1_space X \<and> t1_space Y)"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   222
    using False  
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   223
    apply(simp add: t1_space_closedin_singleton closedin_prod_Times_iff eq 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   224
               del: insert_Times_insert flip: null_topspace_iff_trivial ex_in_conv)
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   225
    by blast
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   226
  with False show ?thesis
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   227
    by simp
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   228
qed
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
   229
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   230
subsection\<open>Hausdorff Spaces\<close>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   231
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   232
definition Hausdorff_space
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   233
  where
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   234
 "Hausdorff_space X \<equiv>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   235
        \<forall>x y. x \<in> topspace X \<and> y \<in> topspace X \<and> (x \<noteq> y)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   236
              \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   237
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   238
lemma Hausdorff_space_expansive:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   239
   "\<lbrakk>Hausdorff_space X; topspace X = topspace Y; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> Hausdorff_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   240
  by (metis Hausdorff_space_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   241
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   242
lemma Hausdorff_space_topspace_empty [iff]: "Hausdorff_space trivial_topology"
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   243
  by (simp add: Hausdorff_space_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   244
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   245
lemma Hausdorff_imp_t1_space:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   246
   "Hausdorff_space X \<Longrightarrow> t1_space X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   247
  by (metis Hausdorff_space_def disjnt_iff t1_space_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   248
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   249
lemma closedin_derived_set_of:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   250
   "Hausdorff_space X \<Longrightarrow> closedin X (X derived_set_of S)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   251
  by (simp add: Hausdorff_imp_t1_space closedin_derived_set_of_gen)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   252
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   253
lemma t1_or_Hausdorff_space:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   254
   "t1_space X \<or> Hausdorff_space X \<longleftrightarrow> t1_space X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   255
  using Hausdorff_imp_t1_space by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   256
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   257
lemma Hausdorff_space_sing_Inter_opens:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   258
   "\<lbrakk>Hausdorff_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> \<Inter>{u. openin X u \<and> a \<in> u} = {a}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   259
  using Hausdorff_imp_t1_space t1_space_singleton_Inter_open by force
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   260
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   261
lemma Hausdorff_space_subtopology:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   262
  assumes "Hausdorff_space X" shows "Hausdorff_space(subtopology X S)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   263
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   264
  have *: "disjnt U V \<Longrightarrow> disjnt (S \<inter> U) (S \<inter> V)" for U V
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   265
    by (simp add: disjnt_iff)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   266
  from assms show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   267
    apply (simp add: Hausdorff_space_def openin_subtopology_alt)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   268
    apply (fast intro: * elim!: all_forward)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   269
    done
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   270
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   271
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   272
lemma Hausdorff_space_compact_separation:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   273
  assumes X: "Hausdorff_space X" and S: "compactin X S" and T: "compactin X T" and "disjnt S T"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   274
  obtains U V where "openin X U" "openin X V" "S \<subseteq> U" "T \<subseteq> V" "disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   275
proof (cases "S = {}")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   276
  case True
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   277
  then show thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   278
    by (metis \<open>compactin X T\<close> compactin_subset_topspace disjnt_empty1 empty_subsetI openin_empty openin_topspace that)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   279
next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   280
  case False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   281
  have "\<forall>x \<in> S. \<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> T \<subseteq> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   282
  proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   283
    fix a
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   284
    assume "a \<in> S"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   285
    then have "a \<notin> T"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   286
      by (meson assms(4) disjnt_iff)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   287
    have a: "a \<in> topspace X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   288
      using S \<open>a \<in> S\<close> compactin_subset_topspace by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   289
    show "\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> T \<subseteq> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   290
    proof (cases "T = {}")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   291
      case True
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   292
      then show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   293
        using a disjnt_empty2 openin_empty by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   294
    next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   295
      case False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   296
      have "\<forall>x \<in> topspace X - {a}. \<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> a \<in> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   297
        using X a by (simp add: Hausdorff_space_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   298
      then obtain U V where UV: "\<forall>x \<in> topspace X - {a}. openin X (U x) \<and> openin X (V x) \<and> x \<in> U x \<and> a \<in> V x \<and> disjnt (U x) (V x)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   299
        by metis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   300
      with \<open>a \<notin> T\<close> compactin_subset_topspace [OF T]
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   301
      have Topen: "\<forall>W \<in> U ` T. openin X W" and Tsub: "T \<subseteq> \<Union> (U ` T)"
75455
91c16c5ad3e9 tidied auto / simp with null arguments
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   302
        by auto
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   303
      then obtain \<F> where \<F>: "finite \<F>" "\<F> \<subseteq> U ` T" and "T \<subseteq> \<Union>\<F>"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   304
        using T unfolding compactin_def by meson
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   305
      then obtain F where F: "finite F" "F \<subseteq> T" "\<F> = U ` F" and SUF: "T \<subseteq> \<Union>(U ` F)" and "a \<notin> F"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   306
        using finite_subset_image [OF \<F>] \<open>a \<notin> T\<close> by (metis subsetD)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   307
      have U: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> openin X (U x)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   308
        and V: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> openin X (V x)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   309
        and disj: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> disjnt (U x) (V x)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   310
        using UV by blast+
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   311
      show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   312
      proof (intro exI conjI)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   313
        have "F \<noteq> {}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   314
          using False SUF by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   315
        with \<open>a \<notin> F\<close> show "openin X (\<Inter>(V ` F))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   316
          using F compactin_subset_topspace [OF T] by (force intro: V)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   317
        show "openin X (\<Union>(U ` F))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   318
          using F Topen Tsub by (force intro: U)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   319
        show "disjnt (\<Inter>(V ` F)) (\<Union>(U ` F))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   320
          using disj
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   321
          apply (auto simp: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   322
          using \<open>F \<subseteq> T\<close> \<open>a \<notin> F\<close> compactin_subset_topspace [OF T] by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   323
        show "a \<in> (\<Inter>(V ` F))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   324
          using \<open>F \<subseteq> T\<close> T UV \<open>a \<notin> T\<close> compactin_subset_topspace by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   325
      qed (auto simp: SUF)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   326
    qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   327
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   328
  then obtain U V where UV: "\<forall>x \<in> S. openin X (U x) \<and> openin X (V x) \<and> x \<in> U x \<and> T \<subseteq> V x \<and> disjnt (U x) (V x)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   329
    by metis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   330
  then have "S \<subseteq> \<Union> (U ` S)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   331
    by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   332
  moreover have "\<forall>W \<in> U ` S. openin X W"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   333
    using UV by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   334
  ultimately obtain I where I: "S \<subseteq> \<Union> (U ` I)" "I \<subseteq> S" "finite I"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   335
    by (metis S compactin_def finite_subset_image)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   336
  show thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   337
  proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   338
    show "openin X (\<Union>(U ` I))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   339
      using \<open>I \<subseteq> S\<close> UV by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   340
    show "openin X (\<Inter> (V ` I))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   341
      using False UV \<open>I \<subseteq> S\<close> \<open>S \<subseteq> \<Union> (U ` I)\<close> \<open>finite I\<close> by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   342
    show "disjnt (\<Union>(U ` I)) (\<Inter> (V ` I))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   343
      by simp (meson UV \<open>I \<subseteq> S\<close> disjnt_subset2 in_mono le_INF_iff order_refl)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   344
  qed (use UV I in auto)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   345
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   346
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   347
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   348
lemma Hausdorff_space_compact_sets:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   349
  "Hausdorff_space X \<longleftrightarrow>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   350
    (\<forall>S T. compactin X S \<and> compactin X T \<and> disjnt S T
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   351
           \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   352
  (is "?lhs = ?rhs")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   353
proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   354
  assume ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   355
  then show ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   356
    by (meson Hausdorff_space_compact_separation)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   357
next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   358
  assume R [rule_format]: ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   359
  show ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   360
  proof (clarsimp simp add: Hausdorff_space_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   361
    fix x y
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   362
    assume "x \<in> topspace X" "y \<in> topspace X" "x \<noteq> y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   363
    then show "\<exists>U. openin X U \<and> (\<exists>V. openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   364
      using R [of "{x}" "{y}"] by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   365
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   366
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   367
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   368
lemma compactin_imp_closedin:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   369
  assumes X: "Hausdorff_space X" and S: "compactin X S" shows "closedin X S"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   370
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   371
  have "S \<subseteq> topspace X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   372
    by (simp add: assms compactin_subset_topspace)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   373
  moreover
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   374
  have "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> topspace X - S" if "x \<in> topspace X" "x \<notin> S" for x
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   375
    using Hausdorff_space_compact_separation [OF X _ S, of "{x}"] that
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   376
    apply (simp add: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   377
    by (metis Diff_mono Diff_triv openin_subset)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   378
  ultimately show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   379
    using closedin_def openin_subopen by force
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   380
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   381
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   382
lemma closedin_Hausdorff_singleton:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   383
   "\<lbrakk>Hausdorff_space X; x \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {x}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   384
  by (simp add: Hausdorff_imp_t1_space closedin_t1_singleton)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   385
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   386
lemma closedin_Hausdorff_sing_eq:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   387
   "Hausdorff_space X \<Longrightarrow> closedin X {x} \<longleftrightarrow> x \<in> topspace X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   388
  by (meson closedin_Hausdorff_singleton closedin_subset insert_subset)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   389
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   390
lemma Hausdorff_space_discrete_topology [simp]:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   391
   "Hausdorff_space (discrete_topology U)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   392
  unfolding Hausdorff_space_def
77935
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   393
  by (metis Hausdorff_space_compact_sets Hausdorff_space_def compactin_discrete_topology equalityE openin_discrete_topology)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   394
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   395
lemma compactin_Int:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   396
   "\<lbrakk>Hausdorff_space X; compactin X S; compactin X T\<rbrakk> \<Longrightarrow> compactin X (S \<inter> T)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   397
  by (simp add: closed_Int_compactin compactin_imp_closedin)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   398
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   399
lemma finite_topspace_imp_discrete_topology:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   400
   "\<lbrakk>topspace X = U; finite U; Hausdorff_space X\<rbrakk> \<Longrightarrow> X = discrete_topology U"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   401
  using Hausdorff_imp_t1_space finite_t1_space_imp_discrete_topology by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   402
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   403
lemma derived_set_of_finite:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   404
   "\<lbrakk>Hausdorff_space X; finite S\<rbrakk> \<Longrightarrow> X derived_set_of S = {}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   405
  using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   406
77935
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   407
lemma infinite_perfect_set:
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   408
   "\<lbrakk>Hausdorff_space X; S \<subseteq> X derived_set_of S; S \<noteq> {}\<rbrakk> \<Longrightarrow> infinite S"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   409
  using derived_set_of_finite by blast
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   410
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   411
lemma derived_set_of_singleton:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   412
   "Hausdorff_space X \<Longrightarrow> X derived_set_of {x} = {}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   413
  by (simp add: derived_set_of_finite)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   414
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   415
lemma closedin_Hausdorff_finite:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   416
   "\<lbrakk>Hausdorff_space X; S \<subseteq> topspace X; finite S\<rbrakk> \<Longrightarrow> closedin X S"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   417
  by (simp add: compactin_imp_closedin finite_imp_compactin_eq)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   418
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   419
lemma open_in_Hausdorff_delete:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   420
   "\<lbrakk>Hausdorff_space X; openin X S\<rbrakk> \<Longrightarrow> openin X (S - {x})"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   421
  using Hausdorff_imp_t1_space t1_space_openin_delete_alt by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   422
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   423
lemma closedin_Hausdorff_finite_eq:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   424
   "\<lbrakk>Hausdorff_space X; finite S\<rbrakk> \<Longrightarrow> closedin X S \<longleftrightarrow> S \<subseteq> topspace X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   425
  by (meson closedin_Hausdorff_finite closedin_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   426
70019
095dce9892e8 A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
   427
lemma derived_set_of_infinite_openin:
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   428
   "Hausdorff_space X
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   429
        \<Longrightarrow> X derived_set_of S =
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   430
            {x \<in> topspace X. \<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite(S \<inter> U)}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   431
  using Hausdorff_imp_t1_space t1_space_derived_set_of_infinite_openin by fastforce
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   432
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   433
lemma Hausdorff_space_discrete_compactin:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   434
   "Hausdorff_space X
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   435
        \<Longrightarrow> S \<inter> X derived_set_of S = {} \<and> compactin X S \<longleftrightarrow> S \<subseteq> topspace X \<and> finite S"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   436
  using derived_set_of_finite discrete_compactin_eq_finite by fastforce
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   437
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   438
lemma Hausdorff_space_finite_topspace:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   439
   "Hausdorff_space X \<Longrightarrow> X derived_set_of (topspace X) = {} \<and> compact_space X \<longleftrightarrow> finite(topspace X)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   440
  using derived_set_of_finite discrete_compact_space_eq_finite by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   441
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   442
lemma derived_set_of_derived_set_subset:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   443
   "Hausdorff_space X \<Longrightarrow> X derived_set_of (X derived_set_of S) \<subseteq> X derived_set_of S"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   444
  by (simp add: Hausdorff_imp_t1_space derived_set_of_derived_set_subset_gen)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   445
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   446
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   447
lemma Hausdorff_space_injective_preimage:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   448
  assumes "Hausdorff_space Y" and cmf: "continuous_map X Y f" and "inj_on f (topspace X)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   449
  shows "Hausdorff_space X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   450
  unfolding Hausdorff_space_def
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   451
proof clarify
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   452
  fix x y
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   453
  assume x: "x \<in> topspace X" and y: "y \<in> topspace X" and "x \<noteq> y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   454
  then obtain U V where "openin Y U" "openin Y V" "f x \<in> U" "f y \<in> V" "disjnt U V"
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   455
    using assms
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   456
    by (smt (verit, ccfv_threshold) Hausdorff_space_def continuous_map image_subset_iff inj_on_def)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   457
  show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   458
  proof (intro exI conjI)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   459
    show "openin X {x \<in> topspace X. f x \<in> U}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   460
      using \<open>openin Y U\<close> cmf continuous_map by fastforce
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   461
    show "openin X {x \<in> topspace X. f x \<in> V}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   462
      using \<open>openin Y V\<close> cmf openin_continuous_map_preimage by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   463
    show "disjnt {x \<in> topspace X. f x \<in> U} {x \<in> topspace X. f x \<in> V}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   464
      using \<open>disjnt U V\<close> by (auto simp add: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   465
  qed (use x \<open>f x \<in> U\<close> y \<open>f y \<in> V\<close> in auto)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   466
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   467
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   468
lemma homeomorphic_Hausdorff_space:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   469
   "X homeomorphic_space Y \<Longrightarrow> Hausdorff_space X \<longleftrightarrow> Hausdorff_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   470
  unfolding homeomorphic_space_def homeomorphic_maps_map
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   471
  by (auto simp: homeomorphic_eq_everything_map Hausdorff_space_injective_preimage)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   472
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   473
lemma Hausdorff_space_retraction_map_image:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   474
   "\<lbrakk>retraction_map X Y r; Hausdorff_space X\<rbrakk> \<Longrightarrow> Hausdorff_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   475
  unfolding retraction_map_def
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   476
  using Hausdorff_space_subtopology homeomorphic_Hausdorff_space retraction_maps_section_image2 by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   477
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   478
lemma compact_Hausdorff_space_optimal:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   479
  assumes eq: "topspace Y = topspace X" and XY: "\<And>U. openin X U \<Longrightarrow> openin Y U"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   480
      and "Hausdorff_space X" "compact_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   481
    shows "Y = X"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   482
proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   483
  have "\<And>U. closedin X U \<Longrightarrow> closedin Y U"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   484
    using XY using topology_finer_closedin [OF eq]
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   485
    by metis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   486
  have "openin Y S = openin X S" for S
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   487
    by (metis XY assms(3) assms(4) closedin_compact_space compactin_contractive compactin_imp_closedin eq openin_closedin_eq)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   488
  then show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   489
    by (simp add: topology_eq)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   490
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   491
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   492
lemma continuous_map_imp_closed_graph:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   493
  assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   494
  shows "closedin (prod_topology X Y) ((\<lambda>x. (x,f x)) ` topspace X)"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   495
  unfolding closedin_def
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   496
proof
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   497
  show "(\<lambda>x. (x, f x)) ` topspace X \<subseteq> topspace (prod_topology X Y)"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   498
    using continuous_map_def f by fastforce
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   499
  show "openin (prod_topology X Y) (topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X)"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   500
    unfolding openin_prod_topology_alt
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   501
  proof (intro allI impI)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   502
    show "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   503
      if "(x,y) \<in> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   504
      for x y
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   505
    proof -
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   506
      have "x \<in> topspace X" and y: "y \<in> topspace Y" "y \<noteq> f x"
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   507
        using that by auto
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   508
      then have "f x \<in> topspace Y"
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   509
        using continuous_map_image_subset_topspace f by blast
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   510
      then obtain U V where UV: "openin Y U" "openin Y V" "f x \<in> U" "y \<in> V" "disjnt U V"
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   511
        using Y y Hausdorff_space_def by metis
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   512
      show ?thesis
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   513
      proof (intro exI conjI)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   514
        show "openin X {x \<in> topspace X. f x \<in> U}"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   515
          using \<open>openin Y U\<close> f openin_continuous_map_preimage by blast
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   516
        show "{x \<in> topspace X. f x \<in> U} \<times> V \<subseteq> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   517
          using UV by (auto simp: disjnt_iff dest: openin_subset)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   518
      qed (use UV \<open>x \<in> topspace X\<close> in auto)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   519
    qed
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   520
  qed
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   521
qed
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   522
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   523
lemma continuous_imp_closed_map:
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   524
   "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y\<rbrakk> \<Longrightarrow> closed_map X Y f"
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   525
  by (meson closed_map_def closedin_compact_space compactin_imp_closedin image_compactin)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   526
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   527
lemma continuous_imp_quotient_map:
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   528
   "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y\<rbrakk>
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   529
        \<Longrightarrow> quotient_map X Y f"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   530
  by (simp add: continuous_imp_closed_map continuous_closed_imp_quotient_map)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   531
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   532
lemma continuous_imp_homeomorphic_map:
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   533
   "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; 
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   534
     f ` (topspace X) = topspace Y; inj_on f (topspace X)\<rbrakk>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   535
        \<Longrightarrow> homeomorphic_map X Y f"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   536
  by (simp add: continuous_imp_closed_map bijective_closed_imp_homeomorphic_map)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   537
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   538
lemma continuous_imp_embedding_map:
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   539
   "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; inj_on f (topspace X)\<rbrakk>
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   540
        \<Longrightarrow> embedding_map X Y f"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   541
  by (simp add: continuous_imp_closed_map injective_closed_imp_embedding_map)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   542
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   543
lemma continuous_inverse_map:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   544
  assumes "compact_space X" "Hausdorff_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   545
    and cmf: "continuous_map X Y f" and gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g(f x) = x"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   546
    and Sf:  "S \<subseteq> f ` (topspace X)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   547
  shows "continuous_map (subtopology Y S) X g"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   548
proof (rule continuous_map_from_subtopology_mono [OF _ \<open>S \<subseteq> f ` (topspace X)\<close>])
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   549
  show "continuous_map (subtopology Y (f ` (topspace X))) X g"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   550
    unfolding continuous_map_closedin
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   551
  proof (intro conjI ballI allI impI)
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   552
    show "g \<in> topspace (subtopology Y (f ` topspace X)) \<rightarrow> topspace X"
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78093
diff changeset
   553
      using gf by auto
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   554
  next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   555
    fix C
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   556
    assume C: "closedin X C"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   557
    show "closedin (subtopology Y (f ` topspace X))
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   558
           {x \<in> topspace (subtopology Y (f ` topspace X)). g x \<in> C}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   559
    proof (rule compactin_imp_closedin)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   560
      show "Hausdorff_space (subtopology Y (f ` topspace X))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   561
        using Hausdorff_space_subtopology [OF \<open>Hausdorff_space Y\<close>] by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   562
      have "compactin Y (f ` C)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   563
        using C cmf image_compactin closedin_compact_space [OF \<open>compact_space X\<close>] by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   564
      moreover have "{x \<in> topspace Y. x \<in> f ` topspace X \<and> g x \<in> C} = f ` C"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   565
        using closedin_subset [OF C] cmf by (auto simp: gf continuous_map_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   566
      ultimately have "compactin Y {x \<in> topspace Y. x \<in> f ` topspace X \<and> g x \<in> C}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   567
        by simp
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   568
      then show "compactin (subtopology Y (f ` topspace X))
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   569
              {x \<in> topspace (subtopology Y (f ` topspace X)). g x \<in> C}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   570
        by (auto simp add: compactin_subtopology)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   571
    qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   572
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   573
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   574
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   575
lemma closed_map_paired_continuous_map_right:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   576
   "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> \<Longrightarrow> closed_map X (prod_topology X Y) (\<lambda>x. (x,f x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   577
  by (simp add: continuous_map_imp_closed_graph embedding_map_graph embedding_imp_closed_map)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   578
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   579
lemma closed_map_paired_continuous_map_left:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   580
  assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   581
  shows "closed_map X (prod_topology Y X) (\<lambda>x. (f x,x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   582
proof -
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   583
  have eq: "(\<lambda>x. (f x,x)) = (\<lambda>(a,b). (b,a)) \<circ> (\<lambda>x. (x,f x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   584
    by auto
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   585
  show ?thesis
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   586
    unfolding eq
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   587
  proof (rule closed_map_compose)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   588
    show "closed_map X (prod_topology X Y) (\<lambda>x. (x, f x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   589
      using Y closed_map_paired_continuous_map_right f by blast
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   590
    show "closed_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(a, b). (b, a))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   591
      by (metis homeomorphic_map_swap homeomorphic_imp_closed_map)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   592
  qed
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   593
qed
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   594
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   595
lemma proper_map_paired_continuous_map_right:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   596
   "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk>
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   597
        \<Longrightarrow> proper_map X (prod_topology X Y) (\<lambda>x. (x,f x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   598
  using closed_injective_imp_proper_map closed_map_paired_continuous_map_right
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   599
  by (metis (mono_tags, lifting) Pair_inject inj_onI)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   600
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   601
lemma proper_map_paired_continuous_map_left:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   602
   "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk>
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   603
        \<Longrightarrow> proper_map X (prod_topology Y X) (\<lambda>x. (f x,x))"
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   604
  using closed_injective_imp_proper_map closed_map_paired_continuous_map_left
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   605
  by (metis (mono_tags, lifting) Pair_inject inj_onI)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70019
diff changeset
   606
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   607
lemma Hausdorff_space_prod_topology:
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   608
  "Hausdorff_space(prod_topology X Y) \<longleftrightarrow> (prod_topology X Y) = trivial_topology \<or> Hausdorff_space X \<and> Hausdorff_space Y"
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   609
  (is "?lhs = ?rhs")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   610
proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   611
  assume ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   612
  then show ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   613
    by (rule topological_property_of_prod_component) (auto simp: Hausdorff_space_subtopology homeomorphic_Hausdorff_space)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   614
next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   615
  assume R: ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   616
  show ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   617
  proof (cases "(topspace X \<times> topspace Y) = {}")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   618
    case False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   619
    with R have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}" and X: "Hausdorff_space X" and Y: "Hausdorff_space Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   620
      by auto
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   621
    show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   622
      unfolding Hausdorff_space_def
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   623
    proof clarify
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   624
      fix x y x' y'
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   625
      assume xy: "(x, y) \<in> topspace (prod_topology X Y)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   626
        and xy': "(x',y') \<in> topspace (prod_topology X Y)"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   627
        and *: "\<nexists>U V. openin (prod_topology X Y) U \<and> openin (prod_topology X Y) V
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   628
               \<and> (x, y) \<in> U \<and> (x', y') \<in> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   629
      have False if "x \<noteq> x' \<or> y \<noteq> y'"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   630
        using that
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   631
      proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   632
        assume "x \<noteq> x'"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   633
        then obtain U V where "openin X U" "openin X V" "x \<in> U" "x' \<in> V" "disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   634
          by (metis Hausdorff_space_def X mem_Sigma_iff topspace_prod_topology xy xy')
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   635
        let ?U = "U \<times> topspace Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   636
        let ?V = "V \<times> topspace Y"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   637
        have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   638
          by (simp_all add: openin_prod_Times_iff \<open>openin X U\<close> \<open>openin X V\<close>)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   639
        moreover have "disjnt ?U ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   640
          by (simp add: \<open>disjnt U V\<close>)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   641
        ultimately show False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   642
          using * \<open>x \<in> U\<close> \<open>x' \<in> V\<close> xy xy' by (metis SigmaD2 SigmaI topspace_prod_topology)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   643
      next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   644
        assume "y \<noteq> y'"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   645
        then obtain U V where "openin Y U" "openin Y V" "y \<in> U" "y' \<in> V" "disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   646
          by (metis Hausdorff_space_def Y mem_Sigma_iff topspace_prod_topology xy xy')
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   647
        let ?U = "topspace X \<times> U"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   648
        let ?V = "topspace X \<times> V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   649
        have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   650
          by (simp_all add: openin_prod_Times_iff \<open>openin Y U\<close> \<open>openin Y V\<close>)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   651
        moreover have "disjnt ?U ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   652
          by (simp add: \<open>disjnt U V\<close>)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   653
        ultimately show False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   654
          using "*" \<open>y \<in> U\<close> \<open>y' \<in> V\<close> xy xy' by (metis SigmaD1 SigmaI topspace_prod_topology)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   655
      qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   656
      then show "x = x' \<and> y = y'"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   657
        by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   658
    qed
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   659
  qed force
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   660
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   661
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   662
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   663
lemma Hausdorff_space_product_topology:
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   664
   "Hausdorff_space (product_topology X I) \<longleftrightarrow> (\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. Hausdorff_space (X i))"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   665
  (is "?lhs = ?rhs")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   666
proof
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   667
  assume ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   668
  then show ?rhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   669
    by (simp add: Hausdorff_space_subtopology PiE_eq_empty_iff homeomorphic_Hausdorff_space 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   670
                  topological_property_of_product_component)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   671
next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   672
  assume R: ?rhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   673
  show ?lhs
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   674
  proof (cases "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {}")
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   675
    case True
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   676
    then show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78320
diff changeset
   677
      by (simp add: Hausdorff_space_def)
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   678
  next
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   679
    case False
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   680
    have "\<exists>U V. openin (product_topology X I) U \<and> openin (product_topology X I) V \<and> f \<in> U \<and> g \<in> V \<and> disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   681
      if f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and g: "g \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and "f \<noteq> g"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   682
      for f g :: "'a \<Rightarrow> 'b"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   683
    proof -
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   684
      obtain m where "f m \<noteq> g m"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   685
        using \<open>f \<noteq> g\<close> by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   686
      then have "m \<in> I"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   687
        using f g by fastforce
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   688
      then have "Hausdorff_space (X m)" 
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   689
        using False that R by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   690
      then obtain U V where U: "openin (X m) U" and V: "openin (X m) V" and "f m \<in> U" "g m \<in> V" "disjnt U V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   691
        by (metis Hausdorff_space_def PiE_mem \<open>f m \<noteq> g m\<close> \<open>m \<in> I\<close> f g)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   692
      show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   693
      proof (intro exI conjI)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   694
        let ?U = "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<inter> {x. x m \<in> U}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   695
        let ?V = "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<inter> {x. x m \<in> V}"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   696
        show "openin (product_topology X I) ?U" "openin (product_topology X I) ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   697
          using \<open>m \<in> I\<close> U V
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   698
          by (force simp add: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)+
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   699
        show "f \<in> ?U"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   700
          using \<open>f m \<in> U\<close> f by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   701
        show "g \<in> ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   702
          using \<open>g m \<in> V\<close> g by blast
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   703
        show "disjnt ?U ?V"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   704
          using \<open>disjnt U V\<close> by (auto simp: PiE_def Pi_def disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   705
        qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   706
    qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   707
    then show ?thesis
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   708
      by (simp add: Hausdorff_space_def)   
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   709
  qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   710
qed
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69986
diff changeset
   711
77935
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   712
lemma Hausdorff_space_closed_neighbourhood:
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   713
   "Hausdorff_space X \<longleftrightarrow>
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   714
    (\<forall>x \<in> topspace X. \<exists>U C. openin X U \<and> closedin X C \<and>
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   715
                      Hausdorff_space(subtopology X C) \<and> x \<in> U \<and> U \<subseteq> C)" (is "_ = ?rhs")
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   716
proof
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   717
  assume R: ?rhs
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   718
  show "Hausdorff_space X"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   719
    unfolding Hausdorff_space_def
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   720
  proof clarify
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   721
    fix x y
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   722
    assume x: "x \<in> topspace X" and y: "y \<in> topspace X" and "x \<noteq> y"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   723
    obtain T C where *: "openin X T" "closedin X C" "x \<in> T" "T \<subseteq> C"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   724
                 and C: "Hausdorff_space (subtopology X C)"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   725
      by (meson R \<open>x \<in> topspace X\<close>)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   726
    show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   727
    proof (cases "y \<in> C")
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   728
      case True
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   729
      with * C obtain U V where U: "openin (subtopology X C) U"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   730
                        and V: "openin (subtopology X C) V"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   731
                        and "x \<in> U" "y \<in> V" "disjnt U V"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   732
        unfolding Hausdorff_space_def
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   733
        by (smt (verit, best) \<open>x \<noteq> y\<close> closedin_subset subsetD topspace_subtopology_subset)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   734
      then obtain U' V' where UV': "U = U' \<inter> C" "openin X U'" "V = V' \<inter> C" "openin X V'"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   735
        by (meson openin_subtopology)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   736
      have "disjnt (T \<inter> U') V'"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   737
        using \<open>disjnt U V\<close> UV' \<open>T \<subseteq> C\<close> by (force simp: disjnt_iff)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   738
      with \<open>T \<subseteq> C\<close> have "disjnt (T \<inter> U') (V' \<union> (topspace X - C))"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   739
        unfolding disjnt_def by blast
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   740
      moreover
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   741
      have "openin X (T \<inter> U')"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   742
        by (simp add: \<open>openin X T\<close> \<open>openin X U'\<close> openin_Int)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   743
      moreover have "openin X (V' \<union> (topspace X - C))"
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   744
        using \<open>closedin X C\<close> \<open>openin X V'\<close> by auto
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   745
      ultimately show ?thesis
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   746
        using UV' \<open>x \<in> T\<close> \<open>x \<in> U\<close> \<open>y \<in> V\<close> by blast
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   747
    next
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   748
      case False
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   749
      with * y show ?thesis
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   750
        by (force simp: closedin_def disjnt_def)
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   751
    qed
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   752
  qed
7f240b0dabd9 More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
   753
qed fastforce
77943
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77935
diff changeset
   754
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
end