src/HOL/Analysis/Abstract_Topology_2.thy
author immler
Mon, 07 Jan 2019 14:06:54 +0100
changeset 69619 3f7d8e05e0f2
parent 69616 d18dc9c5c456
child 69622 003475955593
permissions -rw-r--r--
split off Convex.thy: material that does not require Topology_Euclidean_Space
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     1
(*  Author:     L C Paulson, University of Cambridge
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     3
    Author:     Robert Himmelmann, TU Muenchen
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     4
    Author:     Brian Huffman, Portland State University
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     5
*)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     6
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     7
section \<open>Abstract Topology 2\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     8
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
     9
theory Abstract_Topology_2
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    10
  imports
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    11
    Elementary_Topology
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    12
    Abstract_Topology
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    13
    "HOL-Library.Indicator_Function"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    14
begin
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    15
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    16
text \<open>Combination of Elementary and Abstract Topology\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    17
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    18
(* FIXME: move elsewhere *)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    19
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    20
lemma approachable_lt_le: "(\<exists>(d::real) > 0. \<forall>x. f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> P x)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    21
  apply auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    22
  apply (rule_tac x="d/2" in exI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    23
  apply auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    24
  done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    25
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    26
lemma approachable_lt_le2:  \<comment> \<open>like the above, but pushes aside an extra formula\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    27
    "(\<exists>(d::real) > 0. \<forall>x. Q x \<longrightarrow> f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> Q x \<longrightarrow> P x)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    28
  apply auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    29
  apply (rule_tac x="d/2" in exI, auto)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    30
  done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    31
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    32
lemma triangle_lemma:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    33
  fixes x y z :: real
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    34
  assumes x: "0 \<le> x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    35
    and y: "0 \<le> y"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    36
    and z: "0 \<le> z"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    37
    and xy: "x\<^sup>2 \<le> y\<^sup>2 + z\<^sup>2"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    38
  shows "x \<le> y + z"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    39
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    40
  have "y\<^sup>2 + z\<^sup>2 \<le> y\<^sup>2 + 2 * y * z + z\<^sup>2"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    41
    using z y by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    42
  with xy have th: "x\<^sup>2 \<le> (y + z)\<^sup>2"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    43
    by (simp add: power2_eq_square field_simps)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    44
  from y z have yz: "y + z \<ge> 0"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    45
    by arith
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    46
  from power2_le_imp_le[OF th yz] show ?thesis .
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    47
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    48
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    49
lemma isCont_indicator:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    50
  fixes x :: "'a::t2_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    51
  shows "isCont (indicator A :: 'a \<Rightarrow> real) x = (x \<notin> frontier A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    52
proof auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    53
  fix x
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    54
  assume cts_at: "isCont (indicator A :: 'a \<Rightarrow> real) x" and fr: "x \<in> frontier A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    55
  with continuous_at_open have 1: "\<forall>V::real set. open V \<and> indicator A x \<in> V \<longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    56
    (\<exists>U::'a set. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> V))" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    57
  show False
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    58
  proof (cases "x \<in> A")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    59
    assume x: "x \<in> A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    60
    hence "indicator A x \<in> ({0<..<2} :: real set)" by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    61
    hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({0<..<2} :: real set))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    62
      using 1 open_greaterThanLessThan by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    63
    then guess U .. note U = this
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    64
    hence "\<forall>y\<in>U. indicator A y > (0::real)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    65
      unfolding greaterThanLessThan_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    66
    hence "U \<subseteq> A" using indicator_eq_0_iff by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    67
    hence "x \<in> interior A" using U interiorI by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    68
    thus ?thesis using fr unfolding frontier_def by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    69
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    70
    assume x: "x \<notin> A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    71
    hence "indicator A x \<in> ({-1<..<1} :: real set)" by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    72
    hence "\<exists>U. open U \<and> x \<in> U \<and> (\<forall>y\<in>U. indicator A y \<in> ({-1<..<1} :: real set))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    73
      using 1 open_greaterThanLessThan by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    74
    then guess U .. note U = this
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    75
    hence "\<forall>y\<in>U. indicator A y < (1::real)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    76
      unfolding greaterThanLessThan_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    77
    hence "U \<subseteq> -A" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    78
    hence "x \<in> interior (-A)" using U interiorI by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    79
    thus ?thesis using fr interior_complement unfolding frontier_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    80
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    81
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    82
  assume nfr: "x \<notin> frontier A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    83
  hence "x \<in> interior A \<or> x \<in> interior (-A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    84
    by (auto simp: frontier_def closure_interior)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    85
  thus "isCont ((indicator A)::'a \<Rightarrow> real) x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    86
  proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    87
    assume int: "x \<in> interior A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    88
    then obtain U where U: "open U" "x \<in> U" "U \<subseteq> A" unfolding interior_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    89
    hence "\<forall>y\<in>U. indicator A y = (1::real)" unfolding indicator_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    90
    hence "continuous_on U (indicator A)" by (simp add: continuous_on_const indicator_eq_1_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    91
    thus ?thesis using U continuous_on_eq_continuous_at by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    92
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    93
    assume ext: "x \<in> interior (-A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    94
    then obtain U where U: "open U" "x \<in> U" "U \<subseteq> -A" unfolding interior_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    95
    then have "continuous_on U (indicator A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    96
      using continuous_on_topological by (auto simp: subset_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    97
    thus ?thesis using U continuous_on_eq_continuous_at by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    98
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    99
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   100
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   101
lemma closedin_limpt:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   102
  "closedin (subtopology euclidean T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   103
  apply (simp add: closedin_closed, safe)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   104
   apply (simp add: closed_limpt islimpt_subset)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   105
  apply (rule_tac x="closure S" in exI, simp)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   106
  apply (force simp: closure_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   107
  done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   108
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   109
lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (subtopology euclidean S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   110
  by (meson closedin_limpt closed_subset closedin_closed_trans)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   111
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   112
lemma connected_closed_set:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   113
   "closed S
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   114
    \<Longrightarrow> connected S \<longleftrightarrow> (\<nexists>A B. closed A \<and> closed B \<and> A \<noteq> {} \<and> B \<noteq> {} \<and> A \<union> B = S \<and> A \<inter> B = {})"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   115
  unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   116
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   117
text \<open>If a connnected set is written as the union of two nonempty closed sets, then these sets
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   118
have to intersect.\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   119
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   120
lemma connected_as_closed_union:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   121
  assumes "connected C" "C = A \<union> B" "closed A" "closed B" "A \<noteq> {}" "B \<noteq> {}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   122
  shows "A \<inter> B \<noteq> {}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   123
by (metis assms closed_Un connected_closed_set)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   124
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   125
lemma closedin_subset_trans:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   126
  "closedin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   127
    closedin (subtopology euclidean T) S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   128
  by (meson closedin_limpt subset_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   129
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   130
lemma openin_subset_trans:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   131
  "openin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   132
    openin (subtopology euclidean T) S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   133
  by (auto simp: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   134
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   135
lemma openin_Times:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   136
  "openin (subtopology euclidean S) S' \<Longrightarrow> openin (subtopology euclidean T) T' \<Longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   137
    openin (subtopology euclidean (S \<times> T)) (S' \<times> T')"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   138
  unfolding openin_open using open_Times by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   139
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   140
lemma closedin_compact:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   141
   "\<lbrakk>compact S; closedin (subtopology euclidean S) T\<rbrakk> \<Longrightarrow> compact T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   142
by (metis closedin_closed compact_Int_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   143
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   144
lemma closedin_compact_eq:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   145
  fixes S :: "'a::t2_space set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   146
  shows
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   147
   "compact S
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   148
         \<Longrightarrow> (closedin (subtopology euclidean S) T \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   149
              compact T \<and> T \<subseteq> S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   150
by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   151
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   152
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   153
subsection \<open>Closure\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   154
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   155
lemma closure_openin_Int_closure:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   156
  assumes ope: "openin (subtopology euclidean U) S" and "T \<subseteq> U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   157
  shows "closure(S \<inter> closure T) = closure(S \<inter> T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   158
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   159
  obtain V where "open V" and S: "S = U \<inter> V"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   160
    using ope using openin_open by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   161
  show "closure (S \<inter> closure T) \<subseteq> closure (S \<inter> T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   162
    proof (clarsimp simp: S)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   163
      fix x
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   164
      assume  "x \<in> closure (U \<inter> V \<inter> closure T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   165
      then have "V \<inter> closure T \<subseteq> A \<Longrightarrow> x \<in> closure A" for A
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   166
          by (metis closure_mono subsetD inf.coboundedI2 inf_assoc)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   167
      then have "x \<in> closure (T \<inter> V)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   168
         by (metis \<open>open V\<close> closure_closure inf_commute open_Int_closure_subset)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   169
      then show "x \<in> closure (U \<inter> V \<inter> T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   170
        by (metis \<open>T \<subseteq> U\<close> inf.absorb_iff2 inf_assoc inf_commute)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   171
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   172
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   173
  show "closure (S \<inter> T) \<subseteq> closure (S \<inter> closure T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   174
    by (meson Int_mono closure_mono closure_subset order_refl)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   175
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   176
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   177
corollary infinite_openin:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   178
  fixes S :: "'a :: t1_space set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   179
  shows "\<lbrakk>openin (subtopology euclidean U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   180
  by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   181
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   182
subsection \<open>Frontier\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   183
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   184
lemma connected_Int_frontier:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   185
     "\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   186
  apply (simp add: frontier_interiors connected_openin, safe)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   187
  apply (drule_tac x="s \<inter> interior t" in spec, safe)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   188
   apply (drule_tac [2] x="s \<inter> interior (-t)" in spec)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   189
   apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   190
  done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   191
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   192
subsection \<open>Compactness\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   193
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   194
lemma openin_delete:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   195
  fixes a :: "'a :: t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   196
  shows "openin (subtopology euclidean u) s
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   197
         \<Longrightarrow> openin (subtopology euclidean u) (s - {a})"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   198
by (metis Int_Diff open_delete openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   199
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   200
lemma compact_eq_openin_cover:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   201
  "compact S \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   202
    (\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   203
      (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   204
proof safe
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   205
  fix C
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   206
  assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   207
  then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   208
    unfolding openin_open by force+
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   209
  with \<open>compact S\<close> obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   210
    by (meson compactE)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   211
  then have "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   212
    by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   213
  then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   214
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   215
  assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   216
        (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   217
  show "compact S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   218
  proof (rule compactI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   219
    fix C
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   220
    let ?C = "image (\<lambda>T. S \<inter> T) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   221
    assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   222
    then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   223
      unfolding openin_open by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   224
    with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   225
      by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   226
    let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   227
    have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   228
    proof (intro conjI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   229
      from \<open>D \<subseteq> ?C\<close> show "?D \<subseteq> C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   230
        by (fast intro: inv_into_into)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   231
      from \<open>finite D\<close> show "finite ?D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   232
        by (rule finite_imageI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   233
      from \<open>S \<subseteq> \<Union>D\<close> show "S \<subseteq> \<Union>?D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   234
        apply (rule subset_trans, clarsimp)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   235
        apply (frule subsetD [OF \<open>D \<subseteq> ?C\<close>, THEN f_inv_into_f])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   236
        apply (erule rev_bexI, fast)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   237
        done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   238
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   239
    then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   240
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   241
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   242
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   243
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   244
subsection \<open>Continuity\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   245
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   246
lemma interior_image_subset:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   247
  assumes "inj f" "\<And>x. continuous (at x) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   248
  shows "interior (f ` S) \<subseteq> f ` (interior S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   249
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   250
  fix x assume "x \<in> interior (f ` S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   251
  then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` S" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   252
  then have "x \<in> f ` S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   253
  then obtain y where y: "y \<in> S" "x = f y" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   254
  have "open (f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   255
    using assms \<open>open T\<close> by (simp add: continuous_at_imp_continuous_on open_vimage)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   256
  moreover have "y \<in> vimage f T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   257
    using \<open>x = f y\<close> \<open>x \<in> T\<close> by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   258
  moreover have "vimage f T \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   259
    using \<open>T \<subseteq> image f S\<close> \<open>inj f\<close> unfolding inj_on_def subset_eq by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   260
  ultimately have "y \<in> interior S" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   261
  with \<open>x = f y\<close> show "x \<in> f ` interior S" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   262
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   263
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   264
subsection%unimportant \<open>Equality of continuous functions on closure and related results\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   265
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   266
lemma continuous_closedin_preimage_constant:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   267
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   268
  shows "continuous_on S f \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x = a}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   269
  using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   270
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   271
lemma continuous_closed_preimage_constant:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   272
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   273
  shows "continuous_on S f \<Longrightarrow> closed S \<Longrightarrow> closed {x \<in> S. f x = a}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   274
  using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   275
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   276
lemma continuous_constant_on_closure:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   277
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   278
  assumes "continuous_on (closure S) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   279
      and "\<And>x. x \<in> S \<Longrightarrow> f x = a"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   280
      and "x \<in> closure S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   281
  shows "f x = a"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   282
    using continuous_closed_preimage_constant[of "closure S" f a]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   283
      assms closure_minimal[of S "{x \<in> closure S. f x = a}"] closure_subset
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   284
    unfolding subset_eq
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   285
    by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   286
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   287
lemma image_closure_subset:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   288
  assumes contf: "continuous_on (closure S) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   289
    and "closed T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   290
    and "(f ` S) \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   291
  shows "f ` (closure S) \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   292
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   293
  have "S \<subseteq> {x \<in> closure S. f x \<in> T}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   294
    using assms(3) closure_subset by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   295
  moreover have "closed (closure S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   296
    using continuous_closed_preimage[OF contf] \<open>closed T\<close> by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   297
  ultimately have "closure S = (closure S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   298
    using closure_minimal[of S "(closure S \<inter> f -` T)"] by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   299
  then show ?thesis by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   300
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   301
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   302
subsection%unimportant \<open>A function constant on a set\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   303
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   304
definition constant_on  (infixl "(constant'_on)" 50)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   305
  where "f constant_on A \<equiv> \<exists>y. \<forall>x\<in>A. f x = y"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   306
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   307
lemma constant_on_subset: "\<lbrakk>f constant_on A; B \<subseteq> A\<rbrakk> \<Longrightarrow> f constant_on B"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   308
  unfolding constant_on_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   309
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   310
lemma injective_not_constant:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   311
  fixes S :: "'a::{perfect_space} set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   312
  shows "\<lbrakk>open S; inj_on f S; f constant_on S\<rbrakk> \<Longrightarrow> S = {}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   313
unfolding constant_on_def
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   314
by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   315
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   316
lemma constant_on_closureI:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   317
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   318
  assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   319
    shows "f constant_on (closure S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   320
using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   321
by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   322
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   323
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   324
subsection%unimportant \<open>Continuity relative to a union.\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   325
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   326
lemma continuous_on_Un_local:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   327
    "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t;
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   328
      continuous_on s f; continuous_on t f\<rbrakk>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   329
     \<Longrightarrow> continuous_on (s \<union> t) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   330
  unfolding continuous_on closedin_limpt
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   331
  by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   332
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   333
lemma continuous_on_cases_local:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   334
     "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t;
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   335
       continuous_on s f; continuous_on t g;
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   336
       \<And>x. \<lbrakk>x \<in> s \<and> \<not>P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   337
      \<Longrightarrow> continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   338
  by (rule continuous_on_Un_local) (auto intro: continuous_on_eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   339
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   340
lemma continuous_on_cases_le:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   341
  fixes h :: "'a :: topological_space \<Rightarrow> real"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   342
  assumes "continuous_on {t \<in> s. h t \<le> a} f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   343
      and "continuous_on {t \<in> s. a \<le> h t} g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   344
      and h: "continuous_on s h"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   345
      and "\<And>t. \<lbrakk>t \<in> s; h t = a\<rbrakk> \<Longrightarrow> f t = g t"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   346
    shows "continuous_on s (\<lambda>t. if h t \<le> a then f(t) else g(t))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   347
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   348
  have s: "s = (s \<inter> h -` atMost a) \<union> (s \<inter> h -` atLeast a)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   349
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   350
  have 1: "closedin (subtopology euclidean s) (s \<inter> h -` atMost a)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   351
    by (rule continuous_closedin_preimage [OF h closed_atMost])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   352
  have 2: "closedin (subtopology euclidean s) (s \<inter> h -` atLeast a)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   353
    by (rule continuous_closedin_preimage [OF h closed_atLeast])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   354
  have eq: "s \<inter> h -` {..a} = {t \<in> s. h t \<le> a}" "s \<inter> h -` {a..} = {t \<in> s. a \<le> h t}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   355
    by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   356
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   357
    apply (rule continuous_on_subset [of s, OF _ order_refl])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   358
    apply (subst s)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   359
    apply (rule continuous_on_cases_local)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   360
    using 1 2 s assms apply (auto simp: eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   361
    done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   362
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   363
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   364
lemma continuous_on_cases_1:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   365
  fixes s :: "real set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   366
  assumes "continuous_on {t \<in> s. t \<le> a} f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   367
      and "continuous_on {t \<in> s. a \<le> t} g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   368
      and "a \<in> s \<Longrightarrow> f a = g a"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   369
    shows "continuous_on s (\<lambda>t. if t \<le> a then f(t) else g(t))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   370
using assms
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   371
by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   372
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   373
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   374
subsection%unimportant\<open>Inverse function property for open/closed maps\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   375
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   376
lemma continuous_on_inverse_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   377
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   378
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   379
    and injf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   380
    and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   381
  shows "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   382
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   383
  from imf injf have gTS: "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   384
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   385
  from imf injf have fU: "U \<subseteq> S \<Longrightarrow> (f ` U) = T \<inter> g -` U" for U
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   386
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   387
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   388
    by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   389
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   390
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   391
lemma continuous_on_inverse_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   392
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   393
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   394
    and injf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   395
    and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   396
  shows "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   397
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   398
  from imf injf have gTS: "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   399
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   400
  from imf injf have fU: "U \<subseteq> S \<Longrightarrow> (f ` U) = T \<inter> g -` U" for U
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   401
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   402
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   403
    by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   404
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   405
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   406
lemma homeomorphism_injective_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   407
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   408
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   409
    and injf: "inj_on f S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   410
    and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   411
  obtains g where "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   412
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   413
  have "continuous_on T (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   414
    by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   415
  with imf injf contf show "homeomorphism S T f (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   416
    by (auto simp: homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   417
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   418
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   419
lemma homeomorphism_injective_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   420
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   421
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   422
    and injf: "inj_on f S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   423
    and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   424
  obtains g where "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   425
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   426
  have "continuous_on T (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   427
    by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   428
  with imf injf contf show "homeomorphism S T f (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   429
    by (auto simp: homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   430
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   431
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   432
lemma homeomorphism_imp_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   433
  assumes hom: "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   434
    and oo: "openin (subtopology euclidean S) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   435
  shows "openin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   436
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   437
  from hom oo have [simp]: "f ` U = T \<inter> g -` U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   438
    using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   439
  from hom have "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   440
    unfolding homeomorphism_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   441
  moreover have "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   442
    by (metis hom homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   443
  ultimately show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   444
    by (simp add: continuous_on_open oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   445
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   446
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   447
lemma homeomorphism_imp_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   448
  assumes hom: "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   449
    and oo: "closedin (subtopology euclidean S) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   450
  shows "closedin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   451
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   452
  from hom oo have [simp]: "f ` U = T \<inter> g -` U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   453
    using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   454
  from hom have "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   455
    unfolding homeomorphism_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   456
  moreover have "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   457
    by (metis hom homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   458
  ultimately show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   459
    by (simp add: continuous_on_closed oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   460
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   461
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   462
subsection%unimportant \<open>Seperability\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   463
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   464
lemma subset_second_countable:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   465
  obtains \<B> :: "'a:: second_countable_topology set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   466
    where "countable \<B>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   467
          "{} \<notin> \<B>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   468
          "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   469
          "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   470
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   471
  obtain \<B> :: "'a set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   472
    where "countable \<B>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   473
      and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   474
      and \<B>:    "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   475
  proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   476
    obtain \<C> :: "'a set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   477
      where "countable \<C>" and ope: "\<And>C. C \<in> \<C> \<Longrightarrow> open C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   478
        and \<C>: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<C> \<and> S = \<Union>U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   479
      by (metis univ_second_countable that)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   480
    show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   481
    proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   482
      show "countable ((\<lambda>C. S \<inter> C) ` \<C>)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   483
        by (simp add: \<open>countable \<C>\<close>)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   484
      show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (subtopology euclidean S) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   485
        using ope by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   486
      show "\<And>T. openin (subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   487
        by (metis \<C> image_mono inf_Sup openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   488
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   489
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   490
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   491
  proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   492
    show "countable (\<B> - {{}})"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   493
      using \<open>countable \<B>\<close> by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   494
    show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   495
      by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (subtopology euclidean S) C\<close>)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   496
    show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (subtopology euclidean S) T" for T
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   497
      using \<B> [OF that]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   498
      apply clarify
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   499
      apply (rule_tac x="\<U> - {{}}" in exI, auto)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   500
        done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   501
  qed auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   502
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   503
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   504
lemma Lindelof_openin:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   505
  fixes \<F> :: "'a::second_countable_topology set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   506
  assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean U) S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   507
  obtains \<F>' where "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   508
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   509
  have "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>T. open T \<and> S = U \<inter> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   510
    using assms by (simp add: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   511
  then obtain tf where tf: "\<And>S. S \<in> \<F> \<Longrightarrow> open (tf S) \<and> (S = U \<inter> tf S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   512
    by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   513
  have [simp]: "\<And>\<F>'. \<F>' \<subseteq> \<F> \<Longrightarrow> \<Union>\<F>' = U \<inter> \<Union>(tf ` \<F>')"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   514
    using tf by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   515
  obtain \<G> where "countable \<G> \<and> \<G> \<subseteq> tf ` \<F>" "\<Union>\<G> = \<Union>(tf ` \<F>)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   516
    using tf by (force intro: Lindelof [of "tf ` \<F>"])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   517
  then obtain \<F>' where \<F>': "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   518
    by (clarsimp simp add: countable_subset_image)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   519
  then show ?thesis ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   520
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   521
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   522
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   523
subsection%unimportant\<open>Closed Maps\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   524
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   525
lemma continuous_imp_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   526
  fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   527
  assumes "closedin (subtopology euclidean S) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   528
          "continuous_on S f" "f ` S = T" "compact S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   529
    shows "closedin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   530
  by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   531
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   532
lemma closed_map_restrict:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   533
  assumes cloU: "closedin (subtopology euclidean (S \<inter> f -` T')) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   534
    and cc: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   535
    and "T' \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   536
  shows "closedin (subtopology euclidean T') (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   537
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   538
  obtain V where "closed V" "U = S \<inter> f -` T' \<inter> V"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   539
    using cloU by (auto simp: closedin_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   540
  with cc [of "S \<inter> V"] \<open>T' \<subseteq> T\<close> show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   541
    by (fastforce simp add: closedin_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   542
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   543
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   544
subsection%unimportant\<open>Open Maps\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   545
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   546
lemma open_map_restrict:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   547
  assumes opeU: "openin (subtopology euclidean (S \<inter> f -` T')) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   548
    and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   549
    and "T' \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   550
  shows "openin (subtopology euclidean T') (f ` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   551
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   552
  obtain V where "open V" "U = S \<inter> f -` T' \<inter> V"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   553
    using opeU by (auto simp: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   554
  with oo [of "S \<inter> V"] \<open>T' \<subseteq> T\<close> show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   555
    by (fastforce simp add: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   556
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   557
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   558
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   559
subsection%unimportant\<open>Quotient maps\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   560
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   561
lemma quotient_map_imp_continuous_open:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   562
  assumes T: "f ` S \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   563
      and ope: "\<And>U. U \<subseteq> T
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   564
              \<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   565
                   openin (subtopology euclidean T) U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   566
    shows "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   567
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   568
  have [simp]: "S \<inter> f -` f ` S = S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   569
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   570
    using ope [OF T]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   571
    apply (simp add: continuous_on_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   572
    by (meson ope openin_imp_subset openin_trans)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   573
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   574
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   575
lemma quotient_map_imp_continuous_closed:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   576
  assumes T: "f ` S \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   577
      and ope: "\<And>U. U \<subseteq> T
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   578
                  \<Longrightarrow> (closedin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   579
                       closedin (subtopology euclidean T) U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   580
    shows "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   581
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   582
  have [simp]: "S \<inter> f -` f ` S = S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   583
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   584
    using ope [OF T]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   585
    apply (simp add: continuous_on_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   586
    by (metis (no_types, lifting) ope closedin_imp_subset closedin_trans)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   587
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   588
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   589
lemma open_map_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   590
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   591
      and T: "T \<subseteq> f ` S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   592
      and ope: "\<And>T. openin (subtopology euclidean S) T
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   593
                   \<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   594
    shows "openin (subtopology euclidean S) (S \<inter> f -` T) =
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   595
           openin (subtopology euclidean (f ` S)) T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   596
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   597
  have "T = f ` (S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   598
    using T by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   599
  then show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   600
    using "ope" contf continuous_on_open by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   601
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   602
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   603
lemma closed_map_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   604
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   605
      and T: "T \<subseteq> f ` S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   606
      and ope: "\<And>T. closedin (subtopology euclidean S) T
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   607
              \<Longrightarrow> closedin (subtopology euclidean (f ` S)) (f ` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   608
    shows "openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   609
           openin (subtopology euclidean (f ` S)) T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   610
          (is "?lhs = ?rhs")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   611
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   612
  assume ?lhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   613
  then have *: "closedin (subtopology euclidean S) (S - (S \<inter> f -` T))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   614
    using closedin_diff by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   615
  have [simp]: "(f ` S - f ` (S - (S \<inter> f -` T))) = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   616
    using T by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   617
  show ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   618
    using ope [OF *, unfolded closedin_def] by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   619
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   620
  assume ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   621
  with contf show ?lhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   622
    by (auto simp: continuous_on_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   623
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   624
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   625
lemma continuous_right_inverse_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   626
  assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   627
      and contg: "continuous_on T g" and img: "g ` T \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   628
      and fg [simp]: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   629
      and U: "U \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   630
    shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   631
           openin (subtopology euclidean T) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   632
          (is "?lhs = ?rhs")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   633
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   634
  have f: "\<And>Z. openin (subtopology euclidean (f ` S)) Z \<Longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   635
                openin (subtopology euclidean S) (S \<inter> f -` Z)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   636
  and  g: "\<And>Z. openin (subtopology euclidean (g ` T)) Z \<Longrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   637
                openin (subtopology euclidean T) (T \<inter> g -` Z)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   638
    using contf contg by (auto simp: continuous_on_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   639
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   640
  proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   641
    have "T \<inter> g -` (g ` T \<inter> (S \<inter> f -` U)) = {x \<in> T. f (g x) \<in> U}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   642
      using imf img by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   643
    also have "... = U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   644
      using U by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   645
    finally have eq: "T \<inter> g -` (g ` T \<inter> (S \<inter> f -` U)) = U" .
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   646
    assume ?lhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   647
    then have *: "openin (subtopology euclidean (g ` T)) (g ` T \<inter> (S \<inter> f -` U))"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   648
      by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   649
    show ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   650
      using g [OF *] eq by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   651
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   652
    assume rhs: ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   653
    show ?lhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   654
      by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   655
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   656
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   657
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   658
lemma continuous_left_inverse_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   659
  assumes "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   660
      and "continuous_on (f ` S) g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   661
      and  "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   662
      and "U \<subseteq> f ` S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   663
    shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   664
           openin (subtopology euclidean (f ` S)) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   665
apply (rule continuous_right_inverse_imp_quotient_map)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   666
using assms apply force+
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   667
done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   668
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   669
lemma continuous_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   670
  fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   671
  assumes "continuous_on S f" "f ` S = T" "compact S" "U \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   672
    shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   673
           openin (subtopology euclidean T) U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   674
  by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   675
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   676
subsection%unimportant\<open>Pasting functions together\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   677
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   678
text\<open>on open sets\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   679
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   680
lemma pasting_lemma:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   681
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   682
  assumes clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   683
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   684
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   685
      and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   686
    shows "continuous_on S g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   687
proof (clarsimp simp: continuous_openin_preimage_eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   688
  fix U :: "'b set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   689
  assume "open U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   690
  have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   691
    using clo openin_imp_subset by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   692
  have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   693
    using S f g by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   694
  show "openin (subtopology euclidean S) (S \<inter> g -` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   695
    apply (subst *)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   696
    apply (rule openin_Union, clarify)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   697
    using \<open>open U\<close> clo cont continuous_openin_preimage_gen openin_trans by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   698
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   699
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   700
lemma pasting_lemma_exists:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   701
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   702
  assumes S: "S \<subseteq> (\<Union>i \<in> I. T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   703
      and clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   704
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   705
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   706
    obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   707
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   708
  show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   709
    apply (rule pasting_lemma [OF clo cont])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   710
     apply (blast intro: f)+
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   711
    apply (metis (mono_tags, lifting) S UN_iff subsetCE someI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   712
    done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   713
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   714
  fix x i
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   715
  assume "i \<in> I" "x \<in> S \<inter> T i"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   716
  then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   717
    by (metis (no_types, lifting) IntD2 IntI f someI_ex)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   718
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   719
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   720
text\<open>Likewise on closed sets, with a finiteness assumption\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   721
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   722
lemma pasting_lemma_closed:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   723
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   724
  assumes "finite I"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   725
      and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   726
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   727
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   728
      and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   729
    shows "continuous_on S g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   730
proof (clarsimp simp: continuous_closedin_preimage_eq)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   731
  fix U :: "'b set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   732
  assume "closed U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   733
  have S: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   734
    using clo closedin_imp_subset by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   735
  have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   736
    using S f g by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   737
  show "closedin (subtopology euclidean S) (S \<inter> g -` U)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   738
    apply (subst *)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   739
    apply (rule closedin_Union)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   740
    using \<open>finite I\<close> apply simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   741
    apply (blast intro: \<open>closed U\<close> continuous_closedin_preimage cont clo closedin_trans)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   742
    done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   743
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   744
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   745
lemma pasting_lemma_exists_closed:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   746
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   747
  assumes "finite I"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   748
      and S: "S \<subseteq> (\<Union>i \<in> I. T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   749
      and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   750
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   751
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   752
    obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   753
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   754
  show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   755
    apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   756
     apply (blast intro: f)+
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   757
    apply (metis (mono_tags, lifting) S UN_iff subsetCE someI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   758
    done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   759
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   760
  fix x i
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   761
  assume "i \<in> I" "x \<in> S \<inter> T i"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   762
  then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   763
    by (metis (no_types, lifting) IntD2 IntI f someI_ex)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   764
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   765
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   766
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   767
end