src/HOL/Analysis/Abstract_Topology_2.thy
author paulson <lp15@cam.ac.uk>
Tue, 11 Jul 2023 20:21:58 +0100
changeset 78320 eb9a9690b8f5
parent 78256 71e1aa0d9421
child 78336 6bae28577994
permissions -rw-r--r--
cosmetic improvements, new lemmas, especially more uses of function 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
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
    11
    Elementary_Topology Abstract_Topology Continuum_Not_Denumerable
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    12
    "HOL-Library.Indicator_Function"
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
    13
    "HOL-Library.Equipollence"
69616
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
72225
341b15d092f2 quite a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 71842
diff changeset
    18
lemma approachable_lt_le2: 
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    19
    "(\<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)"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    20
by (meson dense less_eq_real_def order_le_less_trans)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    21
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    22
lemma triangle_lemma:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    23
  fixes x y z :: real
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    24
  assumes x: "0 \<le> x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    25
    and y: "0 \<le> y"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    26
    and z: "0 \<le> z"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    27
    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
    28
  shows "x \<le> y + z"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    29
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    30
  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
    31
    using z y by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    32
  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
    33
    by (simp add: power2_eq_square field_simps)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    34
  from y z have yz: "y + z \<ge> 0"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    35
    by arith
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    36
  from power2_le_imp_le[OF th yz] show ?thesis .
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    37
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    38
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    39
lemma isCont_indicator:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    40
  fixes x :: "'a::t2_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    41
  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
    42
proof auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    43
  fix x
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    44
  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
    45
  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
    46
    (\<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
    47
  show False
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    48
  proof (cases "x \<in> A")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    49
    assume x: "x \<in> A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    50
    hence "indicator A x \<in> ({0<..<2} :: real set)" by simp
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 73932
diff changeset
    51
    with 1 obtain U where U: "open U" "x \<in> U" "\<forall>y\<in>U. indicator A y \<in> ({0<..<2} :: real set)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 73932
diff changeset
    52
      using open_greaterThanLessThan by metis
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    53
    hence "\<forall>y\<in>U. indicator A y > (0::real)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    54
      unfolding greaterThanLessThan_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    55
    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
    56
    hence "x \<in> interior A" using U interiorI by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    57
    thus ?thesis using fr unfolding frontier_def by simp
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    58
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    59
    assume x: "x \<notin> A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    60
    hence "indicator A x \<in> ({-1<..<1} :: real set)" by simp
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 73932
diff changeset
    61
    with 1 obtain U where U: "open U" "x \<in> U" "\<forall>y\<in>U. indicator A y \<in> ({-1<..<1} :: real set)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 73932
diff changeset
    62
      using 1 open_greaterThanLessThan by metis
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    63
    hence "\<forall>y\<in>U. indicator A y < (1::real)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    64
      unfolding greaterThanLessThan_def by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    65
    hence "U \<subseteq> -A" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    66
    hence "x \<in> interior (-A)" using U interiorI by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    67
    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
    68
  qed
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 nfr: "x \<notin> frontier A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    71
  hence "x \<in> interior A \<or> x \<in> interior (-A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    72
    by (auto simp: frontier_def closure_interior)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    73
  thus "isCont ((indicator A)::'a \<Rightarrow> real) x"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    74
  proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    75
    assume int: "x \<in> interior A"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    76
    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
    77
    hence "\<forall>y\<in>U. indicator A y = (1::real)" unfolding indicator_def by auto
71172
nipkow
parents: 70178
diff changeset
    78
    hence "continuous_on U (indicator A)" by (simp add: indicator_eq_1_iff)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    79
    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
    80
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    81
    assume ext: "x \<in> interior (-A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    82
    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
    83
    then have "continuous_on U (indicator A)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    84
      using continuous_on_topological by (auto simp: subset_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    85
    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
    86
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    87
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    88
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    89
lemma islimpt_closure:
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    90
  "\<lbrakk>S \<subseteq> T; \<And>x. \<lbrakk>x islimpt S; x \<in> T\<rbrakk> \<Longrightarrow> x \<in> S\<rbrakk> \<Longrightarrow> S = T \<inter> closure S"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    91
  using closure_def by fastforce
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    92
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
    93
lemma closedin_limpt:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
    94
  "closedin (top_of_set T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> S)"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    95
proof -
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    96
  have "\<And>U x. \<lbrakk>closed U; S = T \<inter> U; x islimpt S; x \<in> T\<rbrakk> \<Longrightarrow> x \<in> S"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    97
    by (meson IntI closed_limpt inf_le2 islimpt_subset)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    98
  then show ?thesis
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    99
    by (metis closed_closure closedin_closed closedin_imp_subset islimpt_closure)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   100
qed
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   101
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   102
lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (top_of_set S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   103
  by (meson closedin_limpt closed_subset closedin_closed_trans)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   104
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   105
lemma connected_closed_set:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   106
   "closed S
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   107
    \<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
   108
  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
   109
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   110
text \<open>If a connnected set is written as the union of two nonempty closed sets, 
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   111
  then these sets have to intersect.\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   112
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   113
lemma connected_as_closed_union:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   114
  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
   115
  shows "A \<inter> B \<noteq> {}"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   116
by (metis assms closed_Un connected_closed_set)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   117
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   118
lemma closedin_subset_trans:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   119
  "closedin (top_of_set U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   120
    closedin (top_of_set T) S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   121
  by (meson closedin_limpt subset_iff)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   122
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   123
lemma openin_subset_trans:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   124
  "openin (top_of_set U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   125
    openin (top_of_set T) S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   126
  by (auto simp: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   127
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   128
lemma closedin_compact:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   129
   "\<lbrakk>compact S; closedin (top_of_set S) T\<rbrakk> \<Longrightarrow> compact T"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   130
by (metis closedin_closed compact_Int_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   131
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   132
lemma closedin_compact_eq:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   133
  fixes S :: "'a::t2_space set"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   134
  shows "compact S \<Longrightarrow> (closedin (top_of_set S) T \<longleftrightarrow> compact T \<and> T \<subseteq> S)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   135
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
   136
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   137
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   138
subsection \<open>Closure\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   139
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   140
lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   141
  by (auto simp: closure_of_def closure_def islimpt_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   142
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   143
lemma closure_openin_Int_closure:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   144
  assumes ope: "openin (top_of_set U) S" and "T \<subseteq> U"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   145
  shows "closure(S \<inter> closure T) = closure(S \<inter> T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   146
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   147
  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
   148
    using ope using openin_open by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   149
  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
   150
    proof (clarsimp simp: S)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   151
      fix x
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   152
      assume  "x \<in> closure (U \<inter> V \<inter> closure T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   153
      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
   154
          by (metis closure_mono subsetD inf.coboundedI2 inf_assoc)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   155
      then have "x \<in> closure (T \<inter> V)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   156
         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
   157
      then show "x \<in> closure (U \<inter> V \<inter> T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   158
        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
   159
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   160
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   161
  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
   162
    by (meson Int_mono closure_mono closure_subset order_refl)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   163
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   164
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   165
corollary infinite_openin:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   166
  fixes S :: "'a :: t1_space set"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   167
  shows "\<lbrakk>openin (top_of_set U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   168
  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
   169
69622
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   170
lemma closure_Int_ballI:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   171
  assumes "\<And>U. \<lbrakk>openin (top_of_set S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
69622
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   172
  shows "S \<subseteq> closure T"
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   173
proof (clarsimp simp: closure_iff_nhds_not_empty)
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   174
  fix x and A and V
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   175
  assume "x \<in> S" "V \<subseteq> A" "open V" "x \<in> V" "T \<inter> A = {}"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   176
  then have "openin (top_of_set S) (A \<inter> V \<inter> S)"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   177
    by (simp add: inf_absorb2 openin_subtopology_Int)
69622
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   178
  moreover have "A \<inter> V \<inter> S \<noteq> {}" using \<open>x \<in> V\<close> \<open>V \<subseteq> A\<close> \<open>x \<in> S\<close>
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   179
    by auto
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   180
  ultimately show False
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   181
    using \<open>T \<inter> A = {}\<close> assms by fastforce
69622
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   182
qed
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   183
003475955593 moved generalized lemmas
immler
parents: 69616
diff changeset
   184
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   185
subsection \<open>Frontier\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   186
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   187
lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   188
  by (auto simp: interior_of_def interior_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   189
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   190
lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   191
  by (auto simp: frontier_of_def frontier_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   192
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   193
lemma connected_Int_frontier:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   194
     "\<lbrakk>connected S; S \<inter> T \<noteq> {}; S - T \<noteq> {}\<rbrakk> \<Longrightarrow> S \<inter> frontier T \<noteq> {}"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   195
  apply (simp add: frontier_interiors connected_openin, safe)
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   196
  apply (drule_tac x="S \<inter> interior T" in spec, safe)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   197
   apply (drule_tac [2] x="S \<inter> interior (-T)" in spec)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   198
   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
   199
  done
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   200
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   201
subsection \<open>Compactness\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   202
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   203
lemma openin_delete:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   204
  fixes a :: "'a :: t1_space"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   205
  shows "openin (top_of_set u) S \<Longrightarrow> openin (top_of_set u) (S - {a})"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   206
by (metis Int_Diff open_delete openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   207
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   208
lemma compact_eq_openin_cover:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   209
  "compact S \<longleftrightarrow>
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   210
    (\<forall>C. (\<forall>c\<in>C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   211
      (\<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
   212
proof safe
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   213
  fix C
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   214
  assume "compact S" and "\<forall>c\<in>C. openin (top_of_set S) c" and "S \<subseteq> \<Union>C"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   215
  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
   216
    unfolding openin_open by force+
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   217
  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
   218
    by (meson compactE)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   219
  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
   220
    by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   221
  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
   222
next
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   223
  assume 1: "\<forall>C. (\<forall>c\<in>C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   224
        (\<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
   225
  show "compact S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   226
  proof (rule compactI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   227
    fix C
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   228
    let ?C = "image (\<lambda>T. S \<inter> T) C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   229
    assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   230
    then have "(\<forall>c\<in>?C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>?C"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   231
      unfolding openin_open by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   232
    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
   233
      by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   234
    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
   235
    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
   236
    proof (intro conjI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   237
      from \<open>D \<subseteq> ?C\<close> show "?D \<subseteq> C"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   238
        by (fast intro: inv_into_into)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   239
      from \<open>finite D\<close> show "finite ?D"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   240
        by (rule finite_imageI)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   241
      from \<open>S \<subseteq> \<Union>D\<close> show "S \<subseteq> \<Union>?D"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   242
        by (metis \<open>D \<subseteq> (\<inter>) S ` C\<close> image_inv_into_cancel inf_Sup le_infE)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   243
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   244
    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
   245
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   246
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   247
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   248
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   249
subsection \<open>Continuity\<close>
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   250
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   251
lemma interior_image_subset:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   252
  assumes "inj f" "\<And>x. continuous (at x) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   253
  shows "interior (f ` S) \<subseteq> f ` (interior S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   254
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   255
  fix x assume "x \<in> interior (f ` S)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   256
  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
   257
  then have "x \<in> f ` S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   258
  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
   259
  have "open (f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   260
    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
   261
  moreover have "y \<in> vimage f T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   262
    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
   263
  moreover have "vimage f T \<subseteq> S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   264
    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
   265
  ultimately have "y \<in> interior S" ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   266
  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
   267
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   268
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   269
subsection\<^marker>\<open>tag unimportant\<close> \<open>Equality of continuous functions on closure and related results\<close>
69616
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_closedin_preimage_constant:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   272
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   273
  shows "continuous_on S f \<Longrightarrow> closedin (top_of_set S) {x \<in> S. f x = a}"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   274
  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
   275
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   276
lemma continuous_closed_preimage_constant:
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
  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
   279
  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
   280
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   281
lemma continuous_constant_on_closure:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   282
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   283
  assumes "continuous_on (closure S) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   284
      and "\<And>x. x \<in> S \<Longrightarrow> f x = a"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   285
      and "x \<in> closure S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   286
  shows "f x = a"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   287
    using continuous_closed_preimage_constant[of "closure S" f a]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   288
      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
   289
    unfolding subset_eq
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   290
    by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   291
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   292
lemma image_closure_subset:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   293
  assumes contf: "continuous_on (closure S) f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   294
    and "closed T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   295
    and "(f ` S) \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   296
  shows "f ` (closure S) \<subseteq> T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   297
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   298
  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
   299
    using assms(3) closure_subset by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   300
  moreover have "closed (closure S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   301
    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
   302
  ultimately have "closure S = (closure S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   303
    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
   304
  then show ?thesis by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   305
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   306
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   307
subsection\<^marker>\<open>tag unimportant\<close> \<open>A function constant on a set\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   308
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   309
definition constant_on  (infixl "(constant'_on)" 50)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   310
  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
   311
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   312
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
   313
  unfolding constant_on_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   314
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   315
lemma injective_not_constant:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   316
  fixes S :: "'a::{perfect_space} set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   317
  shows "\<lbrakk>open S; inj_on f S; f constant_on S\<rbrakk> \<Longrightarrow> S = {}"
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   318
  unfolding constant_on_def
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   319
  by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   320
77138
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   321
lemma constant_on_compose:
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   322
  assumes "f constant_on A"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   323
  shows   "g \<circ> f constant_on A"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   324
  using assms by (auto simp: constant_on_def)
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   325
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   326
lemma not_constant_onI:
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   327
  "f x \<noteq> f y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> \<not>f constant_on A"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   328
  unfolding constant_on_def by metis
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   329
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   330
lemma constant_onE:
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   331
  assumes "f constant_on S" and "\<And>x. x\<in>S \<Longrightarrow> f x = g x"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   332
  shows "g constant_on S"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   333
  using assms unfolding constant_on_def by simp
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   334
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   335
lemma constant_on_closureI:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   336
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   337
  assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f"
77138
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   338
  shows "f constant_on (closure S)"
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   339
  using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def
c8597292cd41 Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 76894
diff changeset
   340
  by metis
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   341
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   342
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   343
subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuity relative to a union.\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   344
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   345
lemma continuous_on_Un_local:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   346
    "\<lbrakk>closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T;
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   347
      continuous_on S f; continuous_on T f\<rbrakk>
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   348
     \<Longrightarrow> continuous_on (S \<union> T) f"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   349
  unfolding continuous_on closedin_limpt
77943
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
   350
  by (metis Lim_trivial_limit Lim_within_Un Un_iff trivial_limit_within)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   351
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   352
lemma continuous_on_cases_local:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   353
     "\<lbrakk>closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T;
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   354
       continuous_on S f; continuous_on T g;
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   355
       \<And>x. \<lbrakk>x \<in> S \<and> \<not>P x \<or> x \<in> T \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk>
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   356
      \<Longrightarrow> continuous_on (S \<union> T) (\<lambda>x. if P x then f x else g x)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   357
  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
   358
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   359
lemma continuous_on_cases_le:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   360
  fixes h :: "'a :: topological_space \<Rightarrow> real"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   361
  assumes "continuous_on {x \<in> S. h x \<le> a} f"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   362
      and "continuous_on {x \<in> S. a \<le> h x} g"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   363
      and h: "continuous_on S h"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   364
      and "\<And>x. \<lbrakk>x \<in> S; h x = a\<rbrakk> \<Longrightarrow> f x = g x"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   365
    shows "continuous_on S (\<lambda>x. if h x \<le> a then f(x) else g(x))"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   366
proof -
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   367
  have S: "S = (S \<inter> h -` atMost a) \<union> (S \<inter> h -` atLeast a)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   368
    by force
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   369
  have 1: "closedin (top_of_set S) (S \<inter> h -` atMost a)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   370
    by (rule continuous_closedin_preimage [OF h closed_atMost])
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   371
  have 2: "closedin (top_of_set S) (S \<inter> h -` atLeast a)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   372
    by (rule continuous_closedin_preimage [OF h closed_atLeast])
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   373
  have [simp]: "S \<inter> h -` {..a} = {x \<in> S. h x \<le> a}" "S \<inter> h -` {a..} = {x \<in> S. a \<le> h x}"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   374
    by auto
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   375
  have "continuous_on (S \<inter> h -` {..a} \<union> S \<inter> h -` {a..}) (\<lambda>x. if h x \<le> a then f x else g x)"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   376
    by (intro continuous_on_cases_local) (use 1 2 S assms in auto)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   377
  then show ?thesis
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   378
    using S by force
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   379
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   380
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   381
lemma continuous_on_cases_1:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   382
  fixes S :: "real set"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   383
  assumes "continuous_on {t \<in> S. t \<le> a} f"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   384
      and "continuous_on {t \<in> S. a \<le> t} g"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   385
      and "a \<in> S \<Longrightarrow> f a = g a"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   386
    shows "continuous_on S (\<lambda>t. if t \<le> a then f(t) else g(t))"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   387
using assms
71172
nipkow
parents: 70178
diff changeset
   388
by (auto intro: continuous_on_cases_le [where h = id, simplified])
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   389
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   390
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   391
subsection\<^marker>\<open>tag unimportant\<close>\<open>Inverse function property for open/closed maps\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   392
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   393
lemma continuous_on_inverse_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   394
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   395
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   396
    and injf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   397
    and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> openin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   398
  shows "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   399
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   400
  from imf injf have gTS: "g ` T = S"
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
  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
   403
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   404
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   405
    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
   406
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   407
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   408
lemma continuous_on_inverse_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   409
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   410
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   411
    and injf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   412
    and oo: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> closedin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   413
  shows "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   414
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   415
  from imf injf have gTS: "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   416
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   417
  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
   418
    by force
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   419
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   420
    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
   421
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   422
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   423
lemma homeomorphism_injective_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   424
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   425
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   426
    and injf: "inj_on f S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   427
    and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> openin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   428
  obtains g where "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   429
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   430
  have "continuous_on T (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   431
    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
   432
  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
   433
    by (auto simp: homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   434
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   435
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   436
lemma homeomorphism_injective_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   437
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   438
    and imf: "f ` S = T"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   439
    and injf: "inj_on f S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   440
    and oo: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> closedin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   441
  obtains g where "homeomorphism S T f g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   442
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   443
  have "continuous_on T (inv_into S f)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   444
    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
   445
  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
   446
    by (auto simp: homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   447
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   448
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   449
lemma homeomorphism_imp_open_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   450
  assumes hom: "homeomorphism S T f g"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   451
    and oo: "openin (top_of_set S) U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   452
  shows "openin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   453
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   454
  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
   455
    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
   456
  from hom have "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   457
    unfolding homeomorphism_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   458
  moreover have "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   459
    by (metis hom homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   460
  ultimately show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   461
    by (simp add: continuous_on_open oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   462
qed
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 homeomorphism_imp_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   465
  assumes hom: "homeomorphism S T f g"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   466
    and oo: "closedin (top_of_set S) U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   467
  shows "closedin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   468
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   469
  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
   470
    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
   471
  from hom have "continuous_on T g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   472
    unfolding homeomorphism_def by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   473
  moreover have "g ` T = S"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   474
    by (metis hom homeomorphism_def)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   475
  ultimately show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   476
    by (simp add: continuous_on_closed oo)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   477
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   478
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   479
subsection\<^marker>\<open>tag unimportant\<close> \<open>Seperability\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   480
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   481
lemma subset_second_countable:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   482
  obtains \<B> :: "'a:: second_countable_topology set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   483
    where "countable \<B>"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   484
          "{} \<notin> \<B>"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   485
          "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   486
          "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   487
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   488
  obtain \<B> :: "'a set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   489
    where "countable \<B>"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   490
      and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   491
      and \<B>:    "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   492
  proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   493
    obtain \<C> :: "'a set set"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   494
      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
   495
        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
   496
      by (metis univ_second_countable that)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   497
    show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   498
    proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   499
      show "countable ((\<lambda>C. S \<inter> C) ` \<C>)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   500
        by (simp add: \<open>countable \<C>\<close>)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   501
      show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (top_of_set S) C"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   502
        using ope by auto
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   503
      show "\<And>T. openin (top_of_set S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   504
        by (metis \<C> image_mono inf_Sup openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   505
    qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   506
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   507
  show ?thesis
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
    show "countable (\<B> - {{}})"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   510
      using \<open>countable \<B>\<close> by blast
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   511
    show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (top_of_set S) C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   512
      by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (top_of_set S) C\<close>)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   513
    show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (top_of_set S) T" for T
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   514
      using \<B> [OF that]
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   515
      apply clarify
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   516
      by (rule_tac x="\<U> - {{}}" in exI, auto)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   517
  qed auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   518
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   519
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   520
lemma Lindelof_openin:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   521
  fixes \<F> :: "'a::second_countable_topology set set"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   522
  assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set U) S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   523
  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
   524
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   525
  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
   526
    using assms by (simp add: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   527
  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
   528
    by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   529
  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
   530
    using tf by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   531
  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
   532
    using tf by (force intro: Lindelof [of "tf ` \<F>"])
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   533
  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
   534
    by (clarsimp simp add: countable_subset_image)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   535
  then show ?thesis ..
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   536
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   537
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   538
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   539
subsection\<^marker>\<open>tag unimportant\<close>\<open>Closed Maps\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   540
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   541
lemma continuous_imp_closed_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   542
  fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   543
  assumes "closedin (top_of_set S) U"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   544
          "continuous_on S f" "f ` S = T" "compact S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   545
    shows "closedin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   546
  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
   547
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   548
lemma closed_map_restrict:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   549
  assumes cloU: "closedin (top_of_set (S \<inter> f -` T')) U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   550
    and cc: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> closedin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   551
    and "T' \<subseteq> T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   552
  shows "closedin (top_of_set T') (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   553
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   554
  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
   555
    using cloU by (auto simp: closedin_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   556
  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
   557
    by (fastforce simp add: closedin_closed)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   558
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   559
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   560
subsection\<^marker>\<open>tag unimportant\<close>\<open>Open Maps\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   561
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   562
lemma open_map_restrict:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   563
  assumes opeU: "openin (top_of_set (S \<inter> f -` T')) U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   564
    and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> openin (top_of_set T) (f ` U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   565
    and "T' \<subseteq> T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   566
  shows "openin (top_of_set T') (f ` U)"
69616
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
  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
   569
    using opeU by (auto simp: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   570
  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
   571
    by (fastforce simp add: openin_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   572
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   573
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   574
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   575
subsection\<^marker>\<open>tag unimportant\<close>\<open>Quotient maps\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   576
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   577
lemma quotient_map_imp_continuous_open:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   578
  assumes T: "f \<in> S \<rightarrow> T"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   579
      and ope: "\<And>U. U \<subseteq> T
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   580
              \<Longrightarrow> (openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   581
                   openin (top_of_set T) U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   582
    shows "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   583
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   584
  have [simp]: "S \<inter> f -` f ` S = S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   585
  show ?thesis
72225
341b15d092f2 quite a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 71842
diff changeset
   586
    by (meson T continuous_on_open_gen ope openin_imp_subset)
69616
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 quotient_map_imp_continuous_closed:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   590
  assumes T: "f \<in> S \<rightarrow> T"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   591
      and ope: "\<And>U. U \<subseteq> T
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   592
                  \<Longrightarrow> (closedin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   593
                       closedin (top_of_set T) U)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   594
    shows "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   595
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   596
  have [simp]: "S \<inter> f -` f ` S = S" by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   597
  show ?thesis
72225
341b15d092f2 quite a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 71842
diff changeset
   598
    by (meson T closedin_imp_subset continuous_on_closed_gen ope)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   599
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   600
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   601
lemma open_map_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   602
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   603
      and T: "T \<subseteq> f ` S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   604
      and ope: "\<And>T. openin (top_of_set S) T
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   605
                   \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   606
    shows "openin (top_of_set S) (S \<inter> f -` T) =
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   607
           openin (top_of_set (f ` S)) T"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   608
proof -
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   609
  have "T = f ` (S \<inter> f -` T)"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   610
    using T by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   611
  then show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   612
    using "ope" contf continuous_on_open by metis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   613
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   614
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   615
lemma closed_map_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   616
  assumes contf: "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   617
      and T: "T \<subseteq> f ` S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   618
      and ope: "\<And>T. closedin (top_of_set S) T
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   619
              \<Longrightarrow> closedin (top_of_set (f ` S)) (f ` T)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   620
    shows "openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   621
           openin (top_of_set (f ` S)) T"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   622
          (is "?lhs = ?rhs")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   623
proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   624
  assume ?lhs
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   625
  then have *: "closedin (top_of_set S) (S - (S \<inter> f -` T))"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   626
    using closedin_diff by fastforce
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   627
  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
   628
    using T by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   629
  show ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   630
    using ope [OF *, unfolded closedin_def] by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   631
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   632
  assume ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   633
  with contf show ?lhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   634
    by (auto simp: continuous_on_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   635
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   636
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   637
lemma continuous_right_inverse_imp_quotient_map:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   638
  assumes contf: "continuous_on S f" and imf: "f \<in> S \<rightarrow> T"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   639
      and contg: "continuous_on T g" and img: "g \<in> T \<rightarrow> S"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   640
      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
   641
      and U: "U \<subseteq> T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   642
    shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   643
           openin (top_of_set T) U"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   644
          (is "?lhs = ?rhs")
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   645
proof -
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   646
  have f: "\<And>Z. openin (top_of_set (f ` S)) Z \<Longrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   647
                openin (top_of_set S) (S \<inter> f -` Z)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   648
  and  g: "\<And>Z. openin (top_of_set (g ` T)) Z \<Longrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   649
                openin (top_of_set T) (T \<inter> g -` Z)"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   650
    using contf contg by (auto simp: continuous_on_open)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   651
  show ?thesis
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   652
  proof
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   653
    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
   654
      using imf img by blast
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   655
    also have "... = U"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   656
      using U by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   657
    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
   658
    assume ?lhs
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   659
    then have *: "openin (top_of_set (g ` T)) (g ` T \<inter> (S \<inter> f -` U))"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   660
      by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self image_subset_iff_funcset)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   661
    show ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   662
      using g [OF *] eq by auto
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   663
  next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   664
    assume rhs: ?rhs
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   665
    show ?lhs
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   666
      using assms continuous_openin_preimage rhs by blast
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   667
  qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   668
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   669
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   670
lemma continuous_left_inverse_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   671
  assumes "continuous_on S f"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   672
      and "continuous_on (f ` S) g"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   673
      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
   674
      and "U \<subseteq> f ` S"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   675
    shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   676
           openin (top_of_set (f ` S)) U"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   677
  using assms 
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   678
  by (intro continuous_right_inverse_imp_quotient_map) auto
69616
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 continuous_imp_quotient_map:
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   681
  fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space"
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   682
  assumes "continuous_on S f" "f ` S = T" "compact S" "U \<subseteq> T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   683
    shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   684
           openin (top_of_set T) U"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   685
  by (simp add: assms closed_map_imp_quotient_map continuous_imp_closed_map)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   686
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   687
subsection\<^marker>\<open>tag unimportant\<close>\<open>Pasting lemmas for functions, for of casewise definitions\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   688
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   689
subsubsection\<open>on open sets\<close>
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   690
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   691
lemma pasting_lemma:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   692
  assumes ope: "\<And>i. i \<in> I \<Longrightarrow> openin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   693
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   694
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   695
      and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   696
    shows "continuous_map X Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   697
  unfolding continuous_map_openin_preimage_eq
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   698
proof (intro conjI allI impI)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   699
  show "g \<in> topspace X \<rightarrow> topspace Y"
71174
nipkow
parents: 71172
diff changeset
   700
    using g cont continuous_map_image_subset_topspace by fastforce
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   701
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   702
  fix U
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   703
  assume Y: "openin Y U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   704
  have T: "T i \<subseteq> topspace X" if "i \<in> I" for i
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   705
    using ope by (simp add: openin_subset that)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   706
  have *: "topspace X \<inter> g -` U = (\<Union>i \<in> I. T i \<inter> f i -` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   707
    using f g T by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   708
  have "\<And>i. i \<in> I \<Longrightarrow> openin X (T i \<inter> f i -` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   709
    using cont unfolding continuous_map_openin_preimage_eq
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   710
    by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   711
  then show "openin X (topspace X \<inter> g -` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   712
    by (auto simp: *)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   713
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   714
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   715
lemma pasting_lemma_exists:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   716
  assumes X: "topspace X \<subseteq> (\<Union>i \<in> I. T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   717
      and ope: "\<And>i. i \<in> I \<Longrightarrow> openin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   718
      and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   719
      and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   720
    obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   721
proof
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   722
  let ?h = "\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   723
  show "continuous_map X Y ?h"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   724
    apply (rule pasting_lemma [OF ope cont])
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   725
     apply (blast intro: f)+
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   726
    by (metis (no_types, lifting) UN_E X subsetD someI_ex)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   727
  show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x" if "i \<in> I" "x \<in> topspace X \<inter> T i" for i x
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   728
    by (metis (no_types, lifting) IntD2 IntI f someI_ex that)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   729
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   730
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   731
lemma pasting_lemma_locally_finite:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   732
  assumes fin: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>V. openin X V \<and> x \<in> V \<and> finite {i \<in> I. T i \<inter> V \<noteq> {}}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   733
    and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   734
    and cont:  "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   735
    and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   736
    and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   737
  shows "continuous_map X Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   738
  unfolding continuous_map_closedin_preimage_eq
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   739
proof (intro conjI allI impI)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   740
  show "g \<in> topspace X \<rightarrow> topspace Y"
71174
nipkow
parents: 71172
diff changeset
   741
    using g cont continuous_map_image_subset_topspace by fastforce
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   742
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   743
  fix U
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   744
  assume Y: "closedin Y U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   745
  have T: "T i \<subseteq> topspace X" if "i \<in> I" for i
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   746
    using clo by (simp add: closedin_subset that)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   747
  have *: "topspace X \<inter> g -` U = (\<Union>i \<in> I. T i \<inter> f i -` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   748
    using f g T by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   749
  have cTf: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i \<inter> f i -` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   750
    using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   751
    by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   752
  have sub: "{Z \<in> (\<lambda>i. T i \<inter> f i -` U) ` I. Z \<inter> V \<noteq> {}}
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   753
           \<subseteq> (\<lambda>i. T i \<inter> f i -` U) ` {i \<in> I. T i \<inter> V \<noteq> {}}" for V
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   754
    by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   755
  have 1: "(\<Union>i\<in>I. T i \<inter> f i -` U) \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   756
    using T by blast
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   757
  then have "locally_finite_in X ((\<lambda>i. T i \<inter> f i -` U) ` I)"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   758
    unfolding locally_finite_in_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   759
    using finite_subset [OF sub] fin by force
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   760
  then show "closedin X (topspace X \<inter> g -` U)"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   761
    by (smt (verit, best) * cTf closedin_locally_finite_Union image_iff)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   762
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   763
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   764
subsubsection\<open>Likewise on closed sets, with a finiteness assumption\<close>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   765
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   766
lemma pasting_lemma_closed:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   767
  assumes fin: "finite I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   768
    and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   769
    and cont:  "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   770
    and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   771
    and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   772
  shows "continuous_map X Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   773
  using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   774
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   775
lemma pasting_lemma_exists_locally_finite:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   776
  assumes fin: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>V. openin X V \<and> x \<in> V \<and> finite {i \<in> I. T i \<inter> V \<noteq> {}}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   777
    and X: "topspace X \<subseteq> \<Union>(T ` I)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   778
    and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   779
    and cont:  "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   780
    and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   781
    and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   782
  obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   783
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   784
  show "continuous_map X Y (\<lambda>x. f(@i. i \<in> I \<and> x \<in> T i) x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   785
    apply (rule pasting_lemma_locally_finite [OF fin])
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   786
        apply (blast intro: assms)+
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   787
    by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   788
next
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   789
  fix x i
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   790
  assume "i \<in> I" and "x \<in> topspace X \<inter> T i"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   791
  then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   792
    by (metis (mono_tags, lifting) IntE IntI f someI2)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   793
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   794
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   795
lemma pasting_lemma_exists_closed:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   796
  assumes fin: "finite I"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   797
    and X: "topspace X \<subseteq> \<Union>(T ` I)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   798
    and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   799
    and cont:  "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   800
    and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   801
  obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   802
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   803
  show "continuous_map X Y (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   804
    apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont])
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   805
     apply (blast intro: f)+
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   806
    by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   807
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   808
  fix x i
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   809
  assume "i \<in> I" "x \<in> topspace X \<inter> T i"
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   810
  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
   811
    by (metis (no_types, lifting) IntD2 IntI f someI_ex)
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   812
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   813
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   814
lemma continuous_map_cases:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   815
  assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   816
      and g: "continuous_map (subtopology X (X closure_of {x. \<not> P x})) Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   817
      and fg: "\<And>x. x \<in> X frontier_of {x. P x} \<Longrightarrow> f x = g x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   818
  shows "continuous_map X Y (\<lambda>x. if P x then f x else g x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   819
proof (rule pasting_lemma_closed)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   820
  let ?f = "\<lambda>b. if b then f else g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   821
  let ?g = "\<lambda>x. if P x then f x else g x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   822
  let ?T = "\<lambda>b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   823
  show "finite {True,False}" by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   824
  have eq: "topspace X - Collect P = topspace X \<inter> {x. \<not> P x}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   825
    by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   826
  show "?f i x = ?f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   827
    if "i \<in> {True,False}" "j \<in> {True,False}" and x: "x \<in> topspace X \<inter> ?T i \<inter> ?T j" for i j x
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   828
  proof -
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   829
    have "f x = g x" if "i" "\<not> j"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   830
      by (smt (verit, best) Diff_Diff_Int closure_of_interior_of closure_of_restrict eq fg 
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   831
          frontier_of_closures interior_of_complement that x)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   832
    moreover
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   833
    have "g x = f x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   834
      if "x \<in> X closure_of {x. \<not> P x}" "x \<in> X closure_of Collect P" "\<not> i" "j" for x
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   835
      by (metis IntI closure_of_restrict eq fg frontier_of_closures that)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   836
    ultimately show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   837
      using that by (auto simp flip: closure_of_restrict)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   838
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   839
  show "\<exists>j. j \<in> {True,False} \<and> x \<in> ?T j \<and> (if P x then f x else g x) = ?f j x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   840
    if "x \<in> topspace X" for x
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   841
    by simp (metis in_closure_of mem_Collect_eq that)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   842
qed (auto simp: f g)
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   843
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   844
lemma continuous_map_cases_alt:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   845
  assumes f: "continuous_map (subtopology X (X closure_of {x \<in> topspace X. P x})) Y f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   846
      and g: "continuous_map (subtopology X (X closure_of {x \<in> topspace X. ~P x})) Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   847
      and fg: "\<And>x. x \<in> X frontier_of {x \<in> topspace X. P x} \<Longrightarrow> f x = g x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   848
    shows "continuous_map X Y (\<lambda>x. if P x then f x else g x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   849
  apply (rule continuous_map_cases)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   850
  using assms
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   851
    apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric])
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   852
  done
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   853
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   854
lemma continuous_map_cases_function:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   855
  assumes contp: "continuous_map X Z p"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   856
    and contf: "continuous_map (subtopology X {x \<in> topspace X. p x \<in> Z closure_of U}) Y f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   857
    and contg: "continuous_map (subtopology X {x \<in> topspace X. p x \<in> Z closure_of (topspace Z - U)}) Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   858
    and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x \<in> Z frontier_of U\<rbrakk> \<Longrightarrow> f x = g x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   859
  shows "continuous_map X Y (\<lambda>x. if p x \<in> U then f x else g x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   860
proof (rule continuous_map_cases_alt)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   861
  show "continuous_map (subtopology X (X closure_of {x \<in> topspace X. p x \<in> U})) Y f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   862
  proof (rule continuous_map_from_subtopology_mono)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   863
    let ?T = "{x \<in> topspace X. p x \<in> Z closure_of U}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   864
    show "continuous_map (subtopology X ?T) Y f"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   865
      by (simp add: contf)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   866
    show "X closure_of {x \<in> topspace X. p x \<in> U} \<subseteq> ?T"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   867
      by (rule continuous_map_closure_preimage_subset [OF contp])
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   868
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   869
  show "continuous_map (subtopology X (X closure_of {x \<in> topspace X. p x \<notin> U})) Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   870
  proof (rule continuous_map_from_subtopology_mono)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   871
    let ?T = "{x \<in> topspace X. p x \<in> Z closure_of (topspace Z - U)}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   872
    show "continuous_map (subtopology X ?T) Y g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   873
      by (simp add: contg)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   874
    have "X closure_of {x \<in> topspace X. p x \<notin> U} \<subseteq> X closure_of {x \<in> topspace X. p x \<in> topspace Z - U}"
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
   875
      by (smt (verit) Collect_mono_iff DiffI closure_of_mono continuous_map contp image_subset_iff)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   876
    then show "X closure_of {x \<in> topspace X. p x \<notin> U} \<subseteq> ?T"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   877
      by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]])
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   878
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   879
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   880
  show "f x = g x" if "x \<in> X frontier_of {x \<in> topspace X. p x \<in> U}" for x
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
   881
    using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   882
qed
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
   883
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   884
subsection \<open>Retractions\<close>
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   885
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   886
definition\<^marker>\<open>tag important\<close> retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   887
where "retraction S T r \<longleftrightarrow>
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   888
  T \<subseteq> S \<and> continuous_on S r \<and> r \<in> S \<rightarrow> T \<and> (\<forall>x\<in>T. r x = x)"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   889
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69939
diff changeset
   890
definition\<^marker>\<open>tag important\<close> retract_of (infixl "retract'_of" 50) where
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   891
"T retract_of S  \<longleftrightarrow>  (\<exists>r. retraction S T r)"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   892
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   893
lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow>  r (r x) = r x"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   894
  unfolding retraction_def by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   895
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   896
text \<open>Preservation of fixpoints under (more general notion of) retraction\<close>
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   897
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   898
lemma invertible_fixpoint_property:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   899
  fixes S :: "'a::topological_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   900
    and T :: "'b::topological_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   901
  assumes contt: "continuous_on T i"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   902
    and "i \<in> T \<rightarrow> S"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   903
    and contr: "continuous_on S r"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   904
    and "r \<in> S \<rightarrow> T"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   905
    and ri: "\<And>y. y \<in> T \<Longrightarrow> r (i y) = y"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   906
    and FP: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   907
    and contg: "continuous_on T g"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   908
    and "g \<in> T \<rightarrow> T"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   909
  obtains y where "y \<in> T" and "g y = y"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   910
proof -
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   911
  have "\<exists>x\<in>S. (i \<circ> g \<circ> r) x = x"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   912
  proof (rule FP)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   913
    show "continuous_on S (i \<circ> g \<circ> r)"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   914
      by (metis assms(4) assms(8) contg continuous_on_compose continuous_on_subset contr contt funcset_image)
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   915
    show "(i \<circ> g \<circ> r) \<in> S \<rightarrow> S"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   916
      using assms(2,4,8) by force
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   917
  qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   918
  then obtain x where x: "x \<in> S" "(i \<circ> g \<circ> r) x = x" ..
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   919
  then have *: "g (r x) \<in> T"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   920
    using assms(4,8) by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   921
  have "r ((i \<circ> g \<circ> r) x) = r x"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   922
    using x by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   923
  then show ?thesis
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   924
    using "*" ri that by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   925
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   926
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   927
lemma homeomorphic_fixpoint_property:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   928
  fixes S :: "'a::topological_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   929
    and T :: "'b::topological_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   930
  assumes "S homeomorphic T"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   931
  shows "(\<forall>f. continuous_on S f \<and> f \<in> S \<rightarrow> S \<longrightarrow> (\<exists>x\<in>S. f x = x)) \<longleftrightarrow>
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   932
         (\<forall>g. continuous_on T g \<and> g \<in> T \<rightarrow> T \<longrightarrow> (\<exists>y\<in>T. g y = y))"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   933
         (is "?lhs = ?rhs")
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   934
proof -
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   935
  obtain r i where r:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   936
      "\<forall>x\<in>S. i (r x) = x" "r ` S = T" "continuous_on S r"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   937
      "\<forall>y\<in>T. r (i y) = y" "i ` T = S" "continuous_on T i"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   938
    using assms unfolding homeomorphic_def homeomorphism_def  by blast
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   939
  show ?thesis
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   940
  proof
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   941
    assume ?lhs
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   942
    with r show ?rhs
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   943
      by (smt (verit, del_insts) Pi_iff image_eqI invertible_fixpoint_property[of T i S r])
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   944
  next
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   945
    assume ?rhs
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   946
    with r show ?lhs
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   947
      by (smt (verit, del_insts) Pi_iff image_eqI invertible_fixpoint_property[of S r T i])
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   948
  qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   949
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   950
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   951
lemma retract_fixpoint_property:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   952
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   953
    and S :: "'a set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   954
  assumes "T retract_of S"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   955
    and FP: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   956
    and contg: "continuous_on T g"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   957
    and "g \<in> T \<rightarrow> T"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   958
  obtains y where "y \<in> T" and "g y = y"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   959
proof -
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   960
  obtain h where "retraction S T h"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   961
    using assms(1) unfolding retract_of_def ..
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   962
  then show ?thesis
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   963
    unfolding retraction_def
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   964
    using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP]
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   965
    by (smt (verit, del_insts) Pi_iff assms(4) contg subsetD that)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   966
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   967
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   968
lemma retraction:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   969
  "retraction S T r \<longleftrightarrow> T \<subseteq> S \<and> continuous_on S r \<and> r ` S = T \<and> (\<forall>x \<in> T. r x = x)"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   970
  by (force simp: retraction_def simp flip: image_subset_iff_funcset)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   971
69753
9a3b4cca6d0b eliminated suspicious Unicode;
wenzelm
parents: 69750
diff changeset
   972
lemma retractionE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close>
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   973
  assumes "retraction S T r"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   974
  obtains "T = r ` S" "r \<in> S \<rightarrow> S" "continuous_on S r" "\<And>x. x \<in> S \<Longrightarrow> r (r x) = r x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   975
proof (rule that)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   976
  from retraction [of S T r] assms
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   977
  have "T \<subseteq> S" "continuous_on S r" "r ` S = T" and "\<forall>x \<in> T. r x = x"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   978
    by simp_all
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   979
  then show  "r \<in> S \<rightarrow> S" "continuous_on S r"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   980
    by auto
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   981
  then show "T = r ` S"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   982
    using \<open>r ` S = T\<close> by blast
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   983
  from \<open>\<forall>x \<in> T. r x = x\<close> have "r x = x" if "x \<in> T" for x
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   984
    using that by simp
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   985
  with \<open>r ` S = T\<close> show "r (r x) = r x" if "x \<in> S" for x
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   986
    using that by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   987
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   988
69753
9a3b4cca6d0b eliminated suspicious Unicode;
wenzelm
parents: 69750
diff changeset
   989
lemma retract_ofE: \<comment> \<open>yields properties normalized wrt. simp -- less likely to loop\<close>
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   990
  assumes "T retract_of S"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
   991
  obtains r where "T = r ` S" "r \<in> S \<rightarrow> S" "continuous_on S r" "\<And>x. x \<in> S \<Longrightarrow> r (r x) = r x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   992
proof -
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   993
  from assms obtain r where "retraction S T r"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   994
    by (auto simp add: retract_of_def)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   995
  with that show thesis
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   996
    by (auto elim: retractionE)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   997
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   998
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
   999
lemma retract_of_imp_extensible:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1000
  assumes "S retract_of T" and "continuous_on S f" and "f \<in> S \<rightarrow> U"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1001
  obtains g where "continuous_on T g" "g \<in> T \<rightarrow> U" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1002
proof -
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1003
  from \<open>S retract_of T\<close> obtain r where "retraction T S r"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1004
    by (auto simp add: retract_of_def)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1005
  then have "continuous_on T (f \<circ> r)"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1006
      by (metis assms(2) continuous_on_compose retraction)
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1007
  then show thesis
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1008
    by (smt (verit, best) Pi_iff \<open>retraction T S r\<close> assms(3) comp_apply retraction_def that)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1009
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1010
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1011
lemma idempotent_imp_retraction:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1012
  assumes "continuous_on S f" and "f \<in> S \<rightarrow> S" and "\<And>x. x \<in> S \<Longrightarrow> f(f x) = f x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1013
    shows "retraction S (f ` S) f"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1014
  by (simp add: assms funcset_image retraction)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1015
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1016
lemma retraction_subset:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1017
  assumes "retraction S T r" and "T \<subseteq> S'" and "S' \<subseteq> S"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1018
  shows "retraction S' T r"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1019
  unfolding retraction_def
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1020
  by (metis assms continuous_on_subset image_mono image_subset_iff_funcset retraction)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1021
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1022
lemma retract_of_subset:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1023
  assumes "T retract_of S" and "T \<subseteq> S'" and "S' \<subseteq> S"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1024
    shows "T retract_of S'"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1025
by (meson assms retract_of_def retraction_subset)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1026
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1027
lemma retraction_refl [simp]: "retraction S S (\<lambda>x. x)"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1028
by (simp add: retraction)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1029
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1030
lemma retract_of_refl [iff]: "S retract_of S"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1031
  unfolding retract_of_def retraction_def
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1032
  using continuous_on_id by blast
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1033
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1034
lemma retract_of_imp_subset:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1035
   "S retract_of T \<Longrightarrow> S \<subseteq> T"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1036
by (simp add: retract_of_def retraction_def)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1037
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1038
lemma retract_of_empty [simp]:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1039
     "({} retract_of S) \<longleftrightarrow> S = {}"  "(S retract_of {}) \<longleftrightarrow> S = {}"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1040
by (auto simp: retract_of_def retraction_def)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1041
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1042
lemma retract_of_singleton [iff]: "({x} retract_of S) \<longleftrightarrow> x \<in> S"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1043
  unfolding retract_of_def retraction_def by force
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1044
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1045
lemma retraction_comp:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1046
   "\<lbrakk>retraction S T f; retraction T U g\<rbrakk> \<Longrightarrow> retraction S U (g \<circ> f)"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1047
  by (smt (verit, best) comp_apply continuous_on_compose image_comp retraction subset_iff)
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1048
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1049
lemma retract_of_trans [trans]:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1050
  assumes "S retract_of T" and "T retract_of U"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1051
    shows "S retract_of U"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1052
using assms by (auto simp: retract_of_def intro: retraction_comp)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1053
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1054
lemma closedin_retract:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1055
  fixes S :: "'a :: t2_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1056
  assumes "S retract_of T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1057
    shows "closedin (top_of_set T) S"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1058
proof -
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1059
  obtain r where r: "S \<subseteq> T" "continuous_on T r" "r \<in> T \<rightarrow> S" "\<And>x. x \<in> S \<Longrightarrow> r x = x"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1060
    using assms by (auto simp: retract_of_def retraction_def)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1061
  have "S = {x\<in>T. x = r x}"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1062
    using r by force
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1063
  also have "\<dots> = T \<inter> ((\<lambda>x. (x, r x)) -` ({y. \<exists>x. y = (x, x)}))"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1064
    unfolding vimage_def mem_Times_iff fst_conv snd_conv
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1065
    using r
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1066
    by auto
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1067
  also have "closedin (top_of_set T) \<dots>"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1068
    by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1069
  finally show ?thesis .
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1070
qed
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1071
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1072
lemma closedin_self [simp]: "closedin (top_of_set S) S"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1073
  by simp
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1074
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1075
lemma retract_of_closed:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1076
    fixes S :: "'a :: t2_space set"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1077
    shows "\<lbrakk>closed T; S retract_of T\<rbrakk> \<Longrightarrow> closed S"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1078
  by (metis closedin_retract closedin_closed_eq)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1079
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1080
lemma retract_of_compact:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1081
     "\<lbrakk>compact T; S retract_of T\<rbrakk> \<Longrightarrow> compact S"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1082
  by (metis compact_continuous_image retract_of_def retraction)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1083
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1084
lemma retract_of_connected:
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1085
    "\<lbrakk>connected T; S retract_of T\<rbrakk> \<Longrightarrow> connected S"
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1086
  by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1087
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1088
lemma retraction_openin_vimage_iff:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1089
  assumes r: "retraction S T r" and "U \<subseteq> T"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1090
  shows "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U" (is "?lhs = ?rhs")
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1091
proof
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1092
  show "?lhs \<Longrightarrow> ?rhs"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1093
    using r
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1094
    by (smt (verit, del_insts) assms(2) continuous_right_inverse_imp_quotient_map image_subset_iff_funcset r retractionE retraction_def retraction_subset)
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1095
  show "?rhs \<Longrightarrow> ?lhs"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1096
    by (metis continuous_on_open r retraction)
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1097
qed
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1098
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1099
lemma retract_of_Times:
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1100
   "\<lbrakk>S retract_of S'; T retract_of T'\<rbrakk> \<Longrightarrow> (S \<times> T) retract_of (S' \<times> T')"
69750
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1101
apply (simp add: retract_of_def retraction_def Sigma_mono, clarify)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1102
apply (rename_tac f g)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1103
apply (rule_tac x="\<lambda>z. ((f \<circ> fst) z, (g \<circ> snd) z)" in exI)
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1104
apply (rule conjI continuous_intros | erule continuous_on_subset | force)+
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1105
done
7d83b0abbfd7 moved generalized material
nipkow
parents: 69622
diff changeset
  1106
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1107
subsection\<open>Retractions on a topological space\<close>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1108
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1109
definition retract_of_space :: "'a set \<Rightarrow> 'a topology \<Rightarrow> bool" (infix "retract'_of'_space" 50)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1110
  where "S retract_of_space X
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1111
         \<equiv> S \<subseteq> topspace X \<and> (\<exists>r. continuous_map X (subtopology X S) r \<and> (\<forall>x \<in> S. r x = x))"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1112
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1113
lemma retract_of_space_retraction_maps:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1114
   "S retract_of_space X \<longleftrightarrow> S \<subseteq> topspace X \<and> (\<exists>r. retraction_maps X (subtopology X S) r id)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1115
  by (auto simp: retract_of_space_def retraction_maps_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1116
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1117
lemma retract_of_space_section_map:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1118
   "S retract_of_space X \<longleftrightarrow> S \<subseteq> topspace X \<and> section_map (subtopology X S) X id"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1119
  unfolding retract_of_space_def retraction_maps_def section_map_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1120
  by (auto simp: continuous_map_from_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1121
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1122
lemma retract_of_space_imp_subset:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1123
   "S retract_of_space X \<Longrightarrow> S \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1124
  by (simp add: retract_of_space_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1125
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1126
lemma retract_of_space_topspace:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1127
   "topspace X retract_of_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1128
  using retract_of_space_def by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1129
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1130
lemma retract_of_space_empty [simp]:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1131
   "{} retract_of_space X \<longleftrightarrow> topspace X = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1132
  by (auto simp: continuous_map_def retract_of_space_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1133
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1134
lemma retract_of_space_singleton [simp]:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1135
  "{a} retract_of_space X \<longleftrightarrow> a \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1136
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1137
  have "continuous_map X (subtopology X {a}) (\<lambda>x. a) \<and> (\<lambda>x. a) a = a" if "a \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1138
    using that by simp
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1139
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1140
    by (force simp: retract_of_space_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1141
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1142
77934
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1143
lemma retract_of_space_trans:
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1144
  assumes "S retract_of_space X"  "T retract_of_space subtopology X S"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1145
  shows "T retract_of_space X"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1146
  using assms
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1147
  apply (simp add: retract_of_space_retraction_maps)
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1148
  by (metis id_comp inf.absorb_iff2 retraction_maps_compose subtopology_subtopology)
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1149
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1150
lemma retract_of_space_subtopology:
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1151
  assumes "S retract_of_space X"  "S \<subseteq> U"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1152
  shows "S retract_of_space subtopology X U"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1153
  using assms
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1154
  apply (clarsimp simp: retract_of_space_def)
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1155
  by (metis continuous_map_from_subtopology inf.absorb2 subtopology_subtopology)
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1156
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1157
lemma retract_of_space_clopen:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1158
  assumes "openin X S" "closedin X S" "S = {} \<Longrightarrow> topspace X = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1159
  shows "S retract_of_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1160
proof (cases "S = {}")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1161
  case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1162
  then obtain a where "a \<in> S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1163
    by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1164
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1165
    unfolding retract_of_space_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1166
  proof (intro exI conjI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1167
    show "S \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1168
      by (simp add: assms closedin_subset)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1169
    have "continuous_map X X (\<lambda>x. if x \<in> S then x else a)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1170
    proof (rule continuous_map_cases)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1171
      show "continuous_map (subtopology X (X closure_of {x. x \<in> S})) X (\<lambda>x. x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1172
        by (simp add: continuous_map_from_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1173
      show "continuous_map (subtopology X (X closure_of {x. x \<notin> S})) X (\<lambda>x. a)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1174
        using \<open>S \<subseteq> topspace X\<close> \<open>a \<in> S\<close> by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1175
      show "x = a" if "x \<in> X frontier_of {x. x \<in> S}" for x
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1176
        using assms that clopenin_eq_frontier_of by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1177
    qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1178
    then show "continuous_map X (subtopology X S) (\<lambda>x. if x \<in> S then x else a)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1179
      using \<open>S \<subseteq> topspace X\<close> \<open>a \<in> S\<close>  by (auto simp: continuous_map_in_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1180
  qed auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1181
qed (use assms in auto)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1182
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1183
lemma retract_of_space_disjoint_union:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1184
  assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \<union> T = topspace X" and "S = {} \<Longrightarrow> topspace X = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1185
  shows "S retract_of_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1186
proof (rule retract_of_space_clopen)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1187
  have "S \<inter> T = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1188
    by (meson ST disjnt_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1189
  then have "S = topspace X - T"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1190
    using ST by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1191
  then show "closedin X S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1192
    using \<open>openin X T\<close> by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1193
qed (auto simp: assms)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1194
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1195
lemma retraction_maps_section_image1:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1196
  assumes "retraction_maps X Y r s"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1197
  shows "s ` (topspace Y) retract_of_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1198
  unfolding retract_of_space_section_map
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1199
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1200
  show "s ` topspace Y \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1201
    using assms continuous_map_image_subset_topspace retraction_maps_def by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1202
  show "section_map (subtopology X (s ` topspace Y)) X id"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1203
    unfolding section_map_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1204
    using assms retraction_maps_to_retract_maps by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1205
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1206
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1207
lemma retraction_maps_section_image2:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1208
   "retraction_maps X Y r s
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1209
        \<Longrightarrow> subtopology X (s ` (topspace Y)) homeomorphic_space Y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1210
  using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1211
        section_map_def by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1212
77934
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1213
lemma hereditary_imp_retractive_property:
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1214
  assumes "\<And>X S. P X \<Longrightarrow> P(subtopology X S)" 
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1215
          "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1216
  assumes "retraction_map X X' r" "P X"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1217
  shows "Q X'"
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1218
  by (meson assms retraction_map_def retraction_maps_section_image2)
01c88cf514fc A few new theorems
paulson <lp15@cam.ac.uk>
parents: 77138
diff changeset
  1219
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1220
subsection\<open>Paths and path-connectedness\<close>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1221
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1222
definition pathin :: "'a topology \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1223
   "pathin X g \<equiv> continuous_map (subtopology euclideanreal {0..1}) X g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1224
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1225
lemma pathin_compose:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1226
     "\<lbrakk>pathin X g; continuous_map X Y f\<rbrakk> \<Longrightarrow> pathin Y (f \<circ> g)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1227
   by (simp add: continuous_map_compose pathin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1228
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1229
lemma pathin_subtopology:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1230
     "pathin (subtopology X S) g \<longleftrightarrow> pathin X g \<and> (\<forall>x \<in> {0..1}. g x \<in> S)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1231
  by (auto simp: pathin_def continuous_map_in_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1232
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1233
lemma pathin_const:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1234
   "pathin X (\<lambda>x. a) \<longleftrightarrow> a \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1235
  by (simp add: pathin_def)
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1236
   
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1237
lemma path_start_in_topspace: "pathin X g \<Longrightarrow> g 0 \<in> topspace X"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1238
  by (force simp: pathin_def continuous_map)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1239
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1240
lemma path_finish_in_topspace: "pathin X g \<Longrightarrow> g 1 \<in> topspace X"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1241
  by (force simp: pathin_def continuous_map)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1242
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1243
lemma path_image_subset_topspace: "pathin X g \<Longrightarrow> g \<in> ({0..1}) \<rightarrow> topspace X"
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1244
  by (force simp: pathin_def continuous_map)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1245
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1246
definition path_connected_space :: "'a topology \<Rightarrow> bool"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1247
  where "path_connected_space X \<equiv> \<forall>x \<in> topspace X. \<forall> y \<in> topspace X. \<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1248
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1249
definition path_connectedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1250
  where "path_connectedin X S \<equiv> S \<subseteq> topspace X \<and> path_connected_space(subtopology X S)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1251
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1252
lemma path_connectedin_absolute [simp]:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1253
     "path_connectedin (subtopology X S) S \<longleftrightarrow> path_connectedin X S"
71172
nipkow
parents: 70178
diff changeset
  1254
  by (simp add: path_connectedin_def subtopology_subtopology)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1255
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1256
lemma path_connectedin_subset_topspace:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1257
     "path_connectedin X S \<Longrightarrow> S \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1258
  by (simp add: path_connectedin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1259
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1260
lemma path_connectedin_subtopology:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1261
     "path_connectedin (subtopology X S) T \<longleftrightarrow> path_connectedin X T \<and> T \<subseteq> S"
71172
nipkow
parents: 70178
diff changeset
  1262
  by (auto simp: path_connectedin_def subtopology_subtopology inf.absorb2)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1263
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1264
lemma path_connectedin:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1265
     "path_connectedin X S \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1266
        S \<subseteq> topspace X \<and>
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1267
        (\<forall>x \<in> S. \<forall>y \<in> S. \<exists>g. pathin X g \<and> g \<in> {0..1} \<rightarrow> S \<and> g 0 = x \<and> g 1 = y)"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1268
  unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1269
  by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2 flip: image_subset_iff_funcset)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1270
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1271
lemma path_connectedin_topspace:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1272
     "path_connectedin X (topspace X) \<longleftrightarrow> path_connected_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1273
  by (simp add: path_connectedin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1274
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1275
lemma path_connected_imp_connected_space:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1276
  assumes "path_connected_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1277
  shows "connected_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1278
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1279
  have *: "\<exists>S. connectedin X S \<and> g 0 \<in> S \<and> g 1 \<in> S" if "pathin X g" for g
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1280
  proof (intro exI conjI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1281
    have "continuous_map (subtopology euclideanreal {0..1}) X g"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1282
      using connectedin_absolute that by (simp add: pathin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1283
    then show "connectedin X (g ` {0..1})"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1284
      by (rule connectedin_continuous_map_image) auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1285
  qed auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1286
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1287
    using assms
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1288
    by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1289
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1290
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1291
lemma path_connectedin_imp_connectedin:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1292
     "path_connectedin X S \<Longrightarrow> connectedin X S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1293
  by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1294
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1295
lemma path_connected_space_topspace_empty:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1296
     "topspace X = {} \<Longrightarrow> path_connected_space X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1297
  by (simp add: path_connected_space_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1298
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1299
lemma path_connectedin_empty [simp]: "path_connectedin X {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1300
  by (simp add: path_connectedin)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1301
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1302
lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \<longleftrightarrow> a \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1303
proof
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1304
  show "path_connectedin X {a} \<Longrightarrow> a \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1305
    by (simp add: path_connectedin)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1306
  show "a \<in> topspace X \<Longrightarrow> path_connectedin X {a}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1307
    unfolding path_connectedin
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1308
    using pathin_const by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1309
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1310
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1311
lemma path_connectedin_continuous_map_image:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1312
  assumes f: "continuous_map X Y f" and S: "path_connectedin X S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1313
  shows "path_connectedin Y (f ` S)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1314
proof -
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1315
  have fX: "f \<in> (topspace X) \<rightarrow> topspace Y"
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1316
    using continuous_map_def f by fastforce
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1317
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1318
    unfolding path_connectedin
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1319
  proof (intro conjI ballI; clarify?)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1320
    fix x
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1321
    assume "x \<in> S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1322
    show "f x \<in> topspace Y"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1323
      using S \<open>x \<in> S\<close> fX path_connectedin_subset_topspace by fastforce
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1324
  next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1325
    fix x y
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1326
    assume "x \<in> S" and "y \<in> S"
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1327
    then obtain g where g: "pathin X g" "g \<in> {0..1} \<rightarrow> S" "g 0 = x" "g 1 = y"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1328
      using S  by (force simp: path_connectedin)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1329
    show "\<exists>g. pathin Y g \<and> g \<in> {0..1} \<rightarrow> f ` S \<and> g 0 = f x \<and> g 1 = f y"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1330
    proof (intro exI conjI)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1331
      show "pathin Y (f \<circ> g)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1332
        using \<open>pathin X g\<close> f pathin_compose by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1333
    qed (use g in auto)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1334
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1335
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1336
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1337
lemma path_connectedin_discrete_topology:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1338
  "path_connectedin (discrete_topology U) S \<longleftrightarrow> S \<subseteq> U \<and> (\<exists>a. S \<subseteq> {a})" (is "?lhs = ?rhs")
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1339
proof
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1340
  show "?lhs \<Longrightarrow> ?rhs"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1341
    by (meson connectedin_discrete_topology path_connectedin_imp_connectedin)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1342
  show "?rhs \<Longrightarrow> ?lhs"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1343
    using subset_singletonD by fastforce
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1344
qed
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1345
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1346
lemma path_connected_space_discrete_topology:
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1347
   "path_connected_space (discrete_topology U) \<longleftrightarrow> (\<exists>a. U \<subseteq> {a})"
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1348
  by (metis path_connectedin_discrete_topology path_connectedin_topspace path_connected_space_topspace_empty
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1349
            subset_singletonD topspace_discrete_topology)
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1350
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1351
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1352
lemma homeomorphic_path_connected_space_imp:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1353
     "\<lbrakk>path_connected_space X; X homeomorphic_space Y\<rbrakk> \<Longrightarrow> path_connected_space Y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1354
  unfolding homeomorphic_space_def homeomorphic_maps_def
78320
eb9a9690b8f5 cosmetic improvements, new lemmas, especially more uses of function space
paulson <lp15@cam.ac.uk>
parents: 78256
diff changeset
  1355
  by (smt (verit, ccfv_threshold) homeomorphic_imp_surjective_map homeomorphic_maps_def homeomorphic_maps_map path_connectedin_continuous_map_image path_connectedin_topspace)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1356
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1357
lemma homeomorphic_path_connected_space:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1358
   "X homeomorphic_space Y \<Longrightarrow> path_connected_space X \<longleftrightarrow> path_connected_space Y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1359
  by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1360
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1361
lemma homeomorphic_map_path_connectedness:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1362
  assumes "homeomorphic_map X Y f" "U \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1363
  shows "path_connectedin Y (f ` U) \<longleftrightarrow> path_connectedin X U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1364
  unfolding path_connectedin_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1365
proof (intro conj_cong homeomorphic_path_connected_space)
78248
740b23f1138a EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents: 78200
diff changeset
  1366
  show "f ` U \<subseteq> topspace Y \<longleftrightarrow> (U \<subseteq> topspace X)"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1367
    using assms homeomorphic_imp_surjective_map by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1368
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1369
  assume "U \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1370
  show "subtopology Y (f ` U) homeomorphic_space subtopology X U"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1371
    using assms unfolding homeomorphic_eq_everything_map
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 72225
diff changeset
  1372
    by (metis (no_types, opaque_lifting) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1373
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1374
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1375
lemma homeomorphic_map_path_connectedness_eq:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1376
   "homeomorphic_map X Y f \<Longrightarrow> path_connectedin X U \<longleftrightarrow> U \<subseteq> topspace X \<and> path_connectedin Y (f ` U)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1377
  by (meson homeomorphic_map_path_connectedness path_connectedin_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1378
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1379
subsection\<open>Connected components\<close>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1380
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1381
definition connected_component_of :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1382
  where "connected_component_of X x y \<equiv>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1383
        \<exists>T. connectedin X T \<and> x \<in> T \<and> y \<in> T"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1384
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1385
abbreviation connected_component_of_set
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1386
  where "connected_component_of_set X x \<equiv> Collect (connected_component_of X x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1387
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1388
definition connected_components_of :: "'a topology \<Rightarrow> ('a set) set"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1389
  where "connected_components_of X \<equiv> connected_component_of_set X ` topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1390
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1391
lemma connected_component_in_topspace:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1392
   "connected_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1393
  by (meson connected_component_of_def connectedin_subset_topspace in_mono)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1394
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1395
lemma connected_component_of_refl:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1396
   "connected_component_of X x x \<longleftrightarrow> x \<in> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1397
  by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1398
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1399
lemma connected_component_of_sym:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1400
   "connected_component_of X x y \<longleftrightarrow> connected_component_of X y x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1401
  by (meson connected_component_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1402
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1403
lemma connected_component_of_trans:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1404
   "\<lbrakk>connected_component_of X x y; connected_component_of X y z\<rbrakk>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1405
        \<Longrightarrow> connected_component_of X x z"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1406
  unfolding connected_component_of_def
71842
db120661dded new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents: 71174
diff changeset
  1407
  using connectedin_Un by blast
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1408
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1409
lemma connected_component_of_mono:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1410
   "\<lbrakk>connected_component_of (subtopology X S) x y; S \<subseteq> T\<rbrakk>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1411
        \<Longrightarrow> connected_component_of (subtopology X T) x y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1412
  by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1413
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1414
lemma connected_component_of_set:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1415
   "connected_component_of_set X x = {y. \<exists>T. connectedin X T \<and> x \<in> T \<and> y \<in> T}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1416
  by (meson connected_component_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1417
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1418
lemma connected_component_of_subset_topspace:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1419
   "connected_component_of_set X x \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1420
  using connected_component_in_topspace by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1421
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1422
lemma connected_component_of_eq_empty:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1423
   "connected_component_of_set X x = {} \<longleftrightarrow> (x \<notin> topspace X)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1424
  using connected_component_in_topspace connected_component_of_refl by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1425
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1426
lemma connected_space_iff_connected_component:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1427
   "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. connected_component_of X x y)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1428
  by (simp add: connected_component_of_def connected_space_subconnected)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1429
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1430
lemma connected_space_imp_connected_component_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1431
   "\<lbrakk>connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1432
    \<Longrightarrow> connected_component_of X a b"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1433
  by (simp add: connected_space_iff_connected_component)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1434
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1435
lemma connected_space_connected_component_set:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1436
   "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. connected_component_of_set X x = topspace X)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1437
  using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1438
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1439
lemma connected_component_of_maximal:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1440
   "\<lbrakk>connectedin X S; x \<in> S\<rbrakk> \<Longrightarrow> S \<subseteq> connected_component_of_set X x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1441
  by (meson Ball_Collect connected_component_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1442
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1443
lemma connected_component_of_equiv:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1444
   "connected_component_of X x y \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1445
    x \<in> topspace X \<and> y \<in> topspace X \<and> connected_component_of X x = connected_component_of X y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1446
  apply (simp add: connected_component_in_topspace fun_eq_iff)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1447
  by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1448
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1449
lemma connected_component_of_disjoint:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1450
   "disjnt (connected_component_of_set X x) (connected_component_of_set X y)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1451
    \<longleftrightarrow> ~(connected_component_of X x y)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1452
  using connected_component_of_equiv unfolding disjnt_iff by force
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1453
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1454
lemma connected_component_of_eq:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1455
   "connected_component_of X x = connected_component_of X y \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1456
        (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1457
        x \<in> topspace X \<and> y \<in> topspace X \<and>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1458
        connected_component_of X x y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1459
  by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1460
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1461
lemma connectedin_connected_component_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1462
   "connectedin X (connected_component_of_set X x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1463
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1464
  have "connected_component_of_set X x = \<Union> {T. connectedin X T \<and> x \<in> T}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1465
    by (auto simp: connected_component_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1466
  then show ?thesis
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1467
    by (metis (no_types, lifting) InterI connectedin_Union emptyE mem_Collect_eq)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1468
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1469
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1470
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1471
lemma Union_connected_components_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1472
   "\<Union>(connected_components_of X) = topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1473
  unfolding connected_components_of_def
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1474
  using connected_component_in_topspace connected_component_of_refl by fastforce
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1475
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1476
lemma connected_components_of_maximal:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1477
   "\<lbrakk>C \<in> connected_components_of X; connectedin X S; ~disjnt C S\<rbrakk> \<Longrightarrow> S \<subseteq> C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1478
  unfolding connected_components_of_def disjnt_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1479
  apply clarify
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1480
  by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1481
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1482
lemma pairwise_disjoint_connected_components_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1483
   "pairwise disjnt (connected_components_of X)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1484
  unfolding connected_components_of_def pairwise_def
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1485
  by (smt (verit, best) connected_component_of_disjoint connected_component_of_eq imageE)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1486
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1487
lemma complement_connected_components_of_Union:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1488
   "C \<in> connected_components_of X
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1489
      \<Longrightarrow> topspace X - C = \<Union> (connected_components_of X - {C})"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1490
  by (metis Union_connected_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1491
      insert_subset pairwise_disjoint_connected_components_of)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1492
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1493
lemma nonempty_connected_components_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1494
   "C \<in> connected_components_of X \<Longrightarrow> C \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1495
  unfolding connected_components_of_def
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1496
  by (metis (no_types, lifting) connected_component_of_eq_empty imageE)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1497
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1498
lemma connected_components_of_subset:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1499
   "C \<in> connected_components_of X \<Longrightarrow> C \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1500
  using Union_connected_components_of by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1501
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1502
lemma connectedin_connected_components_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1503
  assumes "C \<in> connected_components_of X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1504
  shows "connectedin X C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1505
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1506
  have "C \<in> connected_component_of_set X ` topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1507
    using assms connected_components_of_def by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1508
then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1509
  using connectedin_connected_component_of by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1510
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1511
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1512
lemma connected_component_in_connected_components_of:
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1513
  "connected_component_of_set X a \<in> connected_components_of X \<longleftrightarrow> a \<in> topspace X"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1514
  by (metis (no_types, lifting) connected_component_of_eq_empty connected_components_of_def image_iff)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1515
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1516
lemma connected_space_iff_components_eq:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1517
   "connected_space X \<longleftrightarrow> (\<forall>C \<in> connected_components_of X. \<forall>C' \<in> connected_components_of X. C = C')"
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1518
          (is "?lhs = ?rhs")
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1519
proof
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1520
  show "?lhs \<Longrightarrow> ?rhs"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1521
    by (simp add: connected_components_of_def connected_space_connected_component_set)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1522
  show "?rhs \<Longrightarrow> ?lhs"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1523
    by (metis Union_connected_components_of Union_iff connected_space_subconnected connectedin_connected_components_of)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1524
qed
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1525
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1526
lemma connected_components_of_eq_empty:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1527
   "connected_components_of X = {} \<longleftrightarrow> topspace X = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1528
  by (simp add: connected_components_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1529
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1530
lemma connected_components_of_empty_space:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1531
   "topspace X = {} \<Longrightarrow> connected_components_of X = {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1532
  by (simp add: connected_components_of_eq_empty)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1533
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1534
lemma connected_components_of_subset_sing:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1535
   "connected_components_of X \<subseteq> {S} \<longleftrightarrow> connected_space X \<and> (topspace X = {} \<or> topspace X = S)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1536
proof (cases "topspace X = {}")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1537
  case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1538
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1539
    by (simp add: connected_components_of_empty_space connected_space_topspace_empty)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1540
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1541
  case False
76894
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1542
  then have "\<lbrakk>connected_components_of X \<subseteq> {S}\<rbrakk> \<Longrightarrow> topspace X = S"
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1543
    by (metis Sup_empty Union_connected_components_of ccpo_Sup_singleton subset_singleton_iff)
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1544
  with False show ?thesis
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1545
    unfolding connected_components_of_def
23f819af2d9f More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1546
    by (metis connected_space_connected_component_set empty_iff image_subset_iff insert_iff)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1547
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1548
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1549
lemma connected_space_iff_components_subset_singleton:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1550
   "connected_space X \<longleftrightarrow> (\<exists>a. connected_components_of X \<subseteq> {a})"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1551
  by (simp add: connected_components_of_subset_sing)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1552
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1553
lemma connected_components_of_eq_singleton:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1554
   "connected_components_of X = {S}
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1555
\<longleftrightarrow> connected_space X \<and> topspace X \<noteq> {} \<and> S = topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1556
  by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1557
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1558
lemma connected_components_of_connected_space:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1559
   "connected_space X \<Longrightarrow> connected_components_of X = (if topspace X = {} then {} else {topspace X})"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1560
  by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1561
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1562
lemma exists_connected_component_of_superset:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1563
  assumes "connectedin X S" and ne: "topspace X \<noteq> {}"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1564
  shows "\<exists>C. C \<in> connected_components_of X \<and> S \<subseteq> C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1565
proof (cases "S = {}")
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1566
  case True
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1567
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1568
    using ne connected_components_of_def by blast
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1569
next
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1570
  case False
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1571
  then show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1572
    by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1573
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1574
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1575
lemma closedin_connected_components_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1576
  assumes "C \<in> connected_components_of X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1577
  shows   "closedin X C"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1578
proof -
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1579
  obtain x where "x \<in> topspace X" and x: "C = connected_component_of_set X x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1580
    using assms by (auto simp: connected_components_of_def)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1581
  have "connected_component_of_set X x \<subseteq> topspace X"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1582
    by (simp add: connected_component_of_subset_topspace)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1583
  moreover have "X closure_of connected_component_of_set X x \<subseteq> connected_component_of_set X x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1584
  proof (rule connected_component_of_maximal)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1585
    show "connectedin X (X closure_of connected_component_of_set X x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1586
      by (simp add: connectedin_closure_of connectedin_connected_component_of)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1587
    show "x \<in> X closure_of connected_component_of_set X x"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1588
      by (simp add: \<open>x \<in> topspace X\<close> closure_of connected_component_of_refl)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1589
  qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1590
  ultimately
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1591
  show ?thesis
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1592
    using closure_of_subset_eq x by auto
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1593
qed
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1594
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1595
lemma closedin_connected_component_of:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1596
   "closedin X (connected_component_of_set X x)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1597
  by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1598
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1599
lemma connected_component_of_eq_overlap:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1600
   "connected_component_of_set X x = connected_component_of_set X y \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1601
      (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1602
      ~(connected_component_of_set X x \<inter> connected_component_of_set X y = {})"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1603
  using connected_component_of_equiv by fastforce
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1604
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1605
lemma connected_component_of_nonoverlap:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1606
   "connected_component_of_set X x \<inter> connected_component_of_set X y = {} \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1607
     (x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1608
     ~(connected_component_of_set X x = connected_component_of_set X y)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1609
  by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1610
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1611
lemma connected_component_of_overlap:
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1612
   "~(connected_component_of_set X x \<inter> connected_component_of_set X y = {}) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1613
    x \<in> topspace X \<and> y \<in> topspace X \<and>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1614
    connected_component_of_set X x = connected_component_of_set X y"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1615
  by (meson connected_component_of_nonoverlap)
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69753
diff changeset
  1616
77943
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1617
subsection\<open>Combining theorems for continuous functions into the reals\<close>
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1618
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1619
text \<open>The homeomorphism between the real line and the open interval $(-1,1)$\<close>
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1620
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1621
lemma continuous_map_real_shrink:
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1622
  "continuous_map euclideanreal (top_of_set {-1<..<1}) (\<lambda>x. x / (1 + \<bar>x\<bar>))"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1623
proof -
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1624
  have "continuous_on UNIV (\<lambda>x::real. x / (1 + \<bar>x\<bar>))"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1625
    by (intro continuous_intros) auto
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1626
  then show ?thesis
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1627
    by (auto simp: continuous_map_in_subtopology divide_simps)
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1628
qed
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1629
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1630
lemma continuous_on_real_grow:
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1631
  "continuous_on {-1<..<1} (\<lambda>x::real. x / (1 - \<bar>x\<bar>))"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1632
  by (intro continuous_intros) auto
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1633
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1634
lemma real_grow_shrink:
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1635
  fixes x::real 
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1636
  shows "x / (1 + \<bar>x\<bar>) / (1 - \<bar>x / (1 + \<bar>x\<bar>)\<bar>) = x"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1637
  by (force simp: divide_simps)
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1638
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1639
lemma homeomorphic_maps_real_shrink:
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1640
  "homeomorphic_maps euclideanreal (subtopology euclideanreal {-1<..<1}) 
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1641
     (\<lambda>x. x / (1 + \<bar>x\<bar>))  (\<lambda>y. y / (1 - \<bar>y\<bar>))"
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1642
  by (force simp: homeomorphic_maps_def continuous_map_real_shrink continuous_on_real_grow divide_simps)
ffdad62bc235 Importation of additional lemmas from metric.ml
paulson <lp15@cam.ac.uk>
parents: 77934
diff changeset
  1643
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1644
lemma real_shrink_Galois:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1645
  fixes x::real
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1646
  shows "(x / (1 + \<bar>x\<bar>) = y) \<longleftrightarrow> (\<bar>y\<bar> < 1 \<and> y / (1 - \<bar>y\<bar>) = x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1647
  using real_grow_shrink by (fastforce simp add: distrib_left)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1648
78200
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1649
lemma real_shrink_eq:
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1650
  fixes x y::real
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1651
  shows "(x / (1 + \<bar>x\<bar>) = y / (1 + \<bar>y\<bar>)) \<longleftrightarrow> x = y"
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1652
  by (metis real_shrink_Galois)
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1653
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1654
lemma real_shrink_lt:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1655
  fixes x::real
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1656
  shows "x / (1 + \<bar>x\<bar>) < y / (1 + \<bar>y\<bar>) \<longleftrightarrow> x < y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1657
  using zero_less_mult_iff [of x y] by (auto simp: field_simps abs_if not_less)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1658
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1659
lemma real_shrink_le:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1660
  fixes x::real
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1661
  shows "x / (1 + \<bar>x\<bar>) \<le> y / (1 + \<bar>y\<bar>) \<longleftrightarrow> x \<le> y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1662
  by (meson linorder_not_le real_shrink_lt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1663
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1664
lemma real_shrink_grow:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1665
  fixes x::real
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1666
  shows "\<bar>x\<bar> < 1 \<Longrightarrow> x / (1 - \<bar>x\<bar>) / (1 + \<bar>x / (1 - \<bar>x\<bar>)\<bar>) = x"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1667
  using real_shrink_Galois by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1668
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1669
lemma continuous_shrink:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1670
  "continuous_on UNIV (\<lambda>x::real. x / (1 + \<bar>x\<bar>))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1671
  by (intro continuous_intros) auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1672
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1673
lemma strict_mono_shrink:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1674
  "strict_mono (\<lambda>x::real. x / (1 + \<bar>x\<bar>))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1675
  by (simp add: monotoneI real_shrink_lt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1676
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1677
lemma shrink_range: "(\<lambda>x::real. x / (1 + \<bar>x\<bar>)) ` S \<subseteq> {-1<..<1}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1678
  by (auto simp: divide_simps)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1679
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1680
text \<open>Note: connected sets of real numbers are the same thing as intervals\<close>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1681
lemma connected_shrink:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1682
  fixes S :: "real set"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1683
  shows "connected ((\<lambda>x. x / (1 + \<bar>x\<bar>)) ` S) \<longleftrightarrow> connected S"  (is "?lhs = ?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1684
proof 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1685
  assume "?lhs"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1686
  then have "connected ((\<lambda>x. x / (1 - \<bar>x\<bar>)) ` (\<lambda>x. x / (1 + \<bar>x\<bar>)) ` S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1687
    by (metis continuous_on_real_grow shrink_range connected_continuous_image 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1688
               continuous_on_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1689
  then show "?rhs"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1690
    using real_grow_shrink by (force simp add: image_comp)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1691
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1692
  assume ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1693
  then show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1694
    using connected_continuous_image 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1695
    by (metis continuous_on_subset continuous_shrink subset_UNIV)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1696
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 77943
diff changeset
  1697
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1698
subsection \<open>A few cardinality results\<close>
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1699
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1700
lemma eqpoll_real_subset:
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1701
  fixes a b::real
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1702
  assumes "a < b" and S: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> x \<in> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1703
  shows "S \<approx> (UNIV::real set)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1704
proof (rule lepoll_antisym)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1705
  show "S \<lesssim> (UNIV::real set)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1706
    by (simp add: subset_imp_lepoll)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1707
  define f where "f \<equiv> \<lambda>x. (a+b) / 2 + (b-a) / 2 * (x / (1 + \<bar>x\<bar>))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1708
  show "(UNIV::real set) \<lesssim> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1709
    unfolding lepoll_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1710
  proof (intro exI conjI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1711
    show "inj f"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1712
      unfolding inj_on_def f_def
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1713
      by (smt (verit, ccfv_SIG) real_shrink_eq \<open>a<b\<close> divide_eq_0_iff mult_cancel_left times_divide_eq_right)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1714
    have pos: "(b-a) / 2 > 0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1715
      using \<open>a<b\<close> by auto
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1716
    have *: "a < (a + b) / 2 + (b - a) / 2 * x \<longleftrightarrow> (b - a) > (b - a) * -x"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1717
            "(a + b) / 2 + (b - a) / 2 * x < b \<longleftrightarrow> (b - a) * x < (b - a)" for x
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1718
      by (auto simp: field_simps)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1719
    show "range f \<subseteq> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1720
      using shrink_range [of UNIV] \<open>a < b\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1721
      unfolding subset_iff f_def greaterThanLessThan_iff image_iff
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1722
      by (smt (verit, best) S * mult_less_cancel_left2 mult_minus_right)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1723
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1724
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78248
diff changeset
  1725
78256
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1726
lemma reals01_lepoll_nat_sets: "{0..<1::real} \<lesssim> (UNIV::nat set set)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1727
proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1728
  define nxt where "nxt \<equiv> \<lambda>x::real. if x < 1/2 then (True, 2*x) else (False, 2*x - 1)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1729
  have nxt_fun: "nxt \<in> {0..<1} \<rightarrow> UNIV \<times> {0..<1}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1730
    by (simp add: nxt_def Pi_iff)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1731
  define \<sigma> where "\<sigma> \<equiv> \<lambda>x. rec_nat (True, x) (\<lambda>n (b,y). nxt y)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1732
  have \<sigma>Suc [simp]: "\<sigma> x (Suc k) = nxt (snd (\<sigma> x k))" for k x
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1733
    by (simp add: \<sigma>_def case_prod_beta')
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1734
  have \<sigma>01: "x \<in> {0..<1} \<Longrightarrow> \<sigma> x n \<in> UNIV \<times> {0..<1}" for x n
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1735
  proof (induction n)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1736
    case 0
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1737
    then show ?case                                           
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1738
      by (simp add: \<sigma>_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1739
   next
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1740
    case (Suc n)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1741
    with nxt_fun show ?case
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1742
      by (force simp add: Pi_iff split: prod.split)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1743
  qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1744
  define f where "f \<equiv> \<lambda>x. {n. fst (\<sigma> x (Suc n))}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1745
  have snd_nxt: "snd (nxt y) - snd (nxt x) = 2 * (y-x)" 
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1746
    if "fst (nxt x) = fst (nxt y)" for x y
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1747
    using that by (simp add: nxt_def split: if_split_asm)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1748
  have False if "f x = f y" "x < y" "0 \<le> x" "x < 1" "0 \<le> y" "y < 1" for x y :: real
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1749
  proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1750
    have "\<And>k. fst (\<sigma> x (Suc k)) = fst (\<sigma> y (Suc k))"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1751
      using that by (force simp add: f_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1752
    then have eq: "\<And>k. fst (nxt (snd (\<sigma> x k))) = fst (nxt (snd (\<sigma> y k)))"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1753
      by (metis \<sigma>_def case_prod_beta' rec_nat_Suc_imp)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1754
    have *: "snd (\<sigma> y k) - snd (\<sigma> x k) = 2 ^ k * (y-x)" for k
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1755
    proof (induction k)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1756
      case 0
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1757
      then show ?case
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1758
        by (simp add: \<sigma>_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1759
    next
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1760
      case (Suc k)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1761
      then show ?case
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1762
        by (simp add: eq snd_nxt)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1763
    qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1764
    define n where "n \<equiv> nat (\<lceil>log 2 (1 / (y - x))\<rceil>)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1765
    have "2^n \<ge> 1 / (y - x)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1766
      by (simp add: n_def power_of_nat_log_ge)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1767
    then have "2^n * (y-x) \<ge> 1"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1768
      using \<open>x < y\<close> by (simp add: n_def field_simps)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1769
    with * have "snd (\<sigma> y n) - snd (\<sigma> x n) \<ge> 1"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1770
      by presburger
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1771
    moreover have "snd (\<sigma> x n) \<in> {0..<1}" "snd (\<sigma> y n) \<in> {0..<1}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1772
      using that by (meson \<sigma>01 atLeastLessThan_iff mem_Times_iff)+
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1773
    ultimately show False by simp
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1774
  qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1775
  then have "inj_on f {0..<1}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1776
    by (meson atLeastLessThan_iff linorder_inj_onI')
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1777
  then show ?thesis
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1778
    unfolding lepoll_def by blast
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1779
qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1780
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1781
lemma nat_sets_lepoll_reals01: "(UNIV::nat set set) \<lesssim> {0..<1::real}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1782
proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1783
  define F where "F \<equiv> \<lambda>S i. if i\<in>S then (inverse 3::real) ^ i else 0"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1784
  have Fge0: "F S i \<ge> 0" for S i
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1785
    by (simp add: F_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1786
  have F: "summable (F S)" for S
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1787
    unfolding F_def by (force intro: summable_comparison_test_ev [where g = "power (inverse 3)"])
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1788
  have "sum (F S) {..<n} \<le> 3/2" for n S
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1789
  proof (cases n)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1790
    case (Suc n')
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1791
    have "sum (F S) {..<n} \<le> (\<Sum>i<n. inverse 3 ^ i)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1792
      by (simp add: F_def sum_mono)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1793
    also have "\<dots> = (\<Sum>i=0..n'. inverse 3 ^ i)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1794
      using Suc atLeast0AtMost lessThan_Suc_atMost by presburger
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1795
    also have "\<dots> = (3/2) * (1 - inverse 3 ^ n)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1796
      using sum_gp_multiplied [of 0 n' "inverse (3::real)"] by (simp add: Suc field_simps)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1797
    also have "\<dots> \<le> 3/2"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1798
      by (simp add: field_simps)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1799
    finally show ?thesis .
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1800
  qed auto
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1801
  then have F32: "suminf (F S) \<le> 3/2" for S
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1802
    using F suminf_le_const by blast
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1803
  define f where "f \<equiv> \<lambda>S. suminf (F S) / 2"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1804
  have monoF: "F S n \<le> F T n" if "S \<subseteq> T" for S T n
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1805
    using F_def that by auto
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1806
  then have monof: "f S \<le> f T" if "S \<subseteq> T" for S T
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1807
    using that F by (simp add: f_def suminf_le)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1808
  have "f S \<in> {0..<1::real}" for S
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1809
  proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1810
    have "0 \<le> suminf (F S)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1811
      using F by (simp add: F_def suminf_nonneg)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1812
    with F32[of S] show ?thesis
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1813
      by (auto simp: f_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1814
  qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1815
  moreover have "inj f"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1816
  proof
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1817
    fix S T
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1818
    assume "f S = f T" 
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1819
    show "S = T"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1820
    proof (rule ccontr)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1821
      assume "S \<noteq> T"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1822
      then have ST_ne: "(S-T) \<union> (T-S) \<noteq> {}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1823
        by blast
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1824
      define n where "n \<equiv> LEAST n. n \<in> (S-T) \<union> (T-S)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1825
      have sum_split: "suminf (F U) = sum (F U) {..<Suc n} + (\<Sum>k. F U (k + Suc n))"  for U
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1826
        by (metis F add.commute suminf_split_initial_segment)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1827
      have yes: "f U \<ge> (sum (F U) {..<n} + (inverse 3::real) ^ n) / 2" 
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1828
        if "n \<in> U" for U
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1829
      proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1830
        have "0 \<le> (\<Sum>k. F U (k + Suc n))"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1831
          by (metis F Fge0 suminf_nonneg summable_iff_shift)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1832
        moreover have "F U n = (1/3) ^ n"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1833
          by (simp add: F_def that)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1834
        ultimately show ?thesis
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1835
          by (simp add: sum_split f_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1836
      qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1837
      have *: "(\<Sum>k. F UNIV (k + n)) = (\<Sum>k. F UNIV k) * (inverse 3::real) ^ n" for n
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1838
        by (simp add: F_def power_add suminf_mult2)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1839
      have no: "f U < (sum (F U) {..<n} + (inverse 3::real) ^ n) / 2" 
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1840
        if "n \<notin> U" for U
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1841
      proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1842
        have [simp]: "F U n = 0"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1843
          by (simp add: F_def that)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1844
        have "(\<Sum>k. F U (k + Suc n)) \<le> (\<Sum>k. F UNIV (k + Suc n))"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1845
          by (metis F monoF subset_UNIV suminf_le summable_ignore_initial_segment)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1846
        then have "suminf (F U) \<le> (\<Sum>k. F UNIV (k + Suc n)) + (\<Sum>i<n. F U i)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1847
          by (simp add: sum_split)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1848
        also have "\<dots> < (inverse 3::real) ^ n + (\<Sum>i<n. F U i)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1849
          unfolding * using F32[of UNIV] by simp
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1850
        finally have "suminf (F U) < inverse 3 ^ n + sum (F U) {..<n}" .
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1851
        then show ?thesis
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1852
          by (simp add: f_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1853
      qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1854
      have "S \<inter> {..<n} = T \<inter> {..<n}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1855
        using not_less_Least by (fastforce simp add: n_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1856
      then have "sum (F S) {..<n} = sum (F T) {..<n}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1857
        by (metis (no_types, lifting) F_def Int_iff sum.cong)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1858
      moreover consider "n \<in> S-T" | "n \<in> T-S"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1859
        by (metis LeastI_ex ST_ne UnE ex_in_conv n_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1860
      ultimately show False
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1861
        by (smt (verit, best) Diff_iff \<open>f S = f T\<close> yes no)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1862
    qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1863
  qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1864
  ultimately show ?thesis
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1865
    by (meson image_subsetI lepoll_def)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1866
qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1867
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1868
lemma open_interval_eqpoll_reals:
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1869
  fixes a b::real
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1870
  shows "{a<..<b} \<approx> (UNIV::real set) \<longleftrightarrow> a<b"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1871
  using bij_betw_tan bij_betw_open_intervals eqpoll_def
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1872
  by (smt (verit, best) UNIV_I eqpoll_real_subset eqpoll_iff_bijections greaterThanLessThan_iff)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1873
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1874
lemma closed_interval_eqpoll_reals:
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1875
  fixes a b::real
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1876
  shows "{a..b} \<approx> (UNIV::real set) \<longleftrightarrow> a < b"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1877
proof
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1878
  show "{a..b} \<approx> (UNIV::real set) \<Longrightarrow> a < b"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1879
    using eqpoll_finite_iff infinite_Icc_iff infinite_UNIV_char_0 by blast
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1880
qed (auto simp: eqpoll_real_subset)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1881
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1882
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1883
lemma reals_lepoll_reals01: "(UNIV::real set) \<lesssim> {0..<1::real}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1884
proof -
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1885
  have "(UNIV::real set) \<approx> {0<..<1::real}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1886
    by (simp add: open_interval_eqpoll_reals eqpoll_sym)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1887
  also have "\<dots> \<lesssim> {0..<1::real}"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1888
    by (simp add: greaterThanLessThan_subseteq_atLeastLessThan_iff subset_imp_lepoll)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1889
  finally show ?thesis .
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1890
qed
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1891
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1892
lemma nat_sets_eqpoll_reals: "(UNIV::nat set set) \<approx> (UNIV::real set)"
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1893
  by (metis (mono_tags, opaque_lifting) reals_lepoll_reals01 lepoll_antisym lepoll_trans 
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1894
      nat_sets_lepoll_reals01 reals01_lepoll_nat_sets subset_UNIV subset_imp_lepoll)
71e1aa0d9421 A couple of new lemmas involving cardinality
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
  1895
69616
d18dc9c5c456 split off theory combining Elementary_Topology and Abstract_Topology
immler
parents:
diff changeset
  1896
end